author | paulson <lp15@cam.ac.uk> |
Mon, 03 Jul 2023 11:45:59 +0100 | |
changeset 78248 | 740b23f1138a |
parent 74362 | 0135a0c77b64 |
permissions | -rw-r--r-- |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
1 |
(* Title: HOL/Probability/Independent_Family.thy |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
2 |
Author: Johannes Hölzl, TU München |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
3 |
Author: Sudeep Kanav, TU München |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
4 |
*) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
5 |
|
61808 | 6 |
section \<open>Independent families of events, event sets, and random variables\<close> |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
7 |
|
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
8 |
theory Independent_Family |
63626
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents:
63040
diff
changeset
|
9 |
imports Infinite_Product_Measure |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
10 |
begin |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
11 |
|
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
12 |
definition (in prob_space) |
42983 | 13 |
"indep_sets F I \<longleftrightarrow> (\<forall>i\<in>I. F i \<subseteq> events) \<and> |
42981 | 14 |
(\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> (\<forall>A\<in>Pi J F. prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))))" |
15 |
||
16 |
definition (in prob_space) |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
17 |
"indep_set A B \<longleftrightarrow> indep_sets (case_bool A B) UNIV" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
18 |
|
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
19 |
definition (in prob_space) |
49784
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
20 |
indep_events_def_alt: "indep_events A I \<longleftrightarrow> indep_sets (\<lambda>i. {A i}) I" |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
21 |
|
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
22 |
lemma (in prob_space) indep_events_def: |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
23 |
"indep_events A I \<longleftrightarrow> (A`I \<subseteq> events) \<and> |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
24 |
(\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j)))" |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
25 |
unfolding indep_events_def_alt indep_sets_def |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
26 |
apply (simp add: Ball_def Pi_iff image_subset_iff_funcset) |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
27 |
apply (intro conj_cong refl arg_cong[where f=All] ext imp_cong) |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
28 |
apply auto |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
29 |
done |
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
30 |
|
59000 | 31 |
lemma (in prob_space) indep_eventsI: |
32 |
"(\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M) \<Longrightarrow> (\<And>J. J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> J \<noteq> {} \<Longrightarrow> prob (\<Inter>i\<in>J. F i) = (\<Prod>i\<in>J. prob (F i))) \<Longrightarrow> indep_events F I" |
|
33 |
by (auto simp: indep_events_def) |
|
34 |
||
49784
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
35 |
definition (in prob_space) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
36 |
"indep_event A B \<longleftrightarrow> indep_events (case_bool A B) UNIV" |
49784
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
37 |
|
47694 | 38 |
lemma (in prob_space) indep_sets_cong: |
42981 | 39 |
"I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i = G i) \<Longrightarrow> indep_sets F I \<longleftrightarrow> indep_sets G J" |
40 |
by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+ |
|
41 |
||
42 |
lemma (in prob_space) indep_events_finite_index_events: |
|
43 |
"indep_events F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_events F J)" |
|
44 |
by (auto simp: indep_events_def) |
|
45 |
||
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
46 |
lemma (in prob_space) indep_sets_finite_index_sets: |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
47 |
"indep_sets F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_sets F J)" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
48 |
proof (intro iffI allI impI) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
49 |
assume *: "\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_sets F J" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
50 |
show "indep_sets F I" unfolding indep_sets_def |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
51 |
proof (intro conjI ballI allI impI) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
52 |
fix i assume "i \<in> I" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
53 |
with *[THEN spec, of "{i}"] show "F i \<subseteq> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
54 |
by (auto simp: indep_sets_def) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
55 |
qed (insert *, auto simp: indep_sets_def) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
56 |
qed (auto simp: indep_sets_def) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
57 |
|
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
58 |
lemma (in prob_space) indep_sets_mono_index: |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
59 |
"J \<subseteq> I \<Longrightarrow> indep_sets F I \<Longrightarrow> indep_sets F J" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
60 |
unfolding indep_sets_def by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
61 |
|
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
62 |
lemma (in prob_space) indep_sets_mono_sets: |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
63 |
assumes indep: "indep_sets F I" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
64 |
assumes mono: "\<And>i. i\<in>I \<Longrightarrow> G i \<subseteq> F i" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
65 |
shows "indep_sets G I" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
66 |
proof - |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
67 |
have "(\<forall>i\<in>I. F i \<subseteq> events) \<Longrightarrow> (\<forall>i\<in>I. G i \<subseteq> events)" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
68 |
using mono by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
69 |
moreover have "\<And>A J. J \<subseteq> I \<Longrightarrow> A \<in> (\<Pi> j\<in>J. G j) \<Longrightarrow> A \<in> (\<Pi> j\<in>J. F j)" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
70 |
using mono by (auto simp: Pi_iff) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
71 |
ultimately show ?thesis |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
72 |
using indep by (auto simp: indep_sets_def) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
73 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
74 |
|
49772 | 75 |
lemma (in prob_space) indep_sets_mono: |
76 |
assumes indep: "indep_sets F I" |
|
77 |
assumes mono: "J \<subseteq> I" "\<And>i. i\<in>J \<Longrightarrow> G i \<subseteq> F i" |
|
78 |
shows "indep_sets G J" |
|
79 |
apply (rule indep_sets_mono_sets) |
|
80 |
apply (rule indep_sets_mono_index) |
|
81 |
apply (fact +) |
|
82 |
done |
|
83 |
||
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
84 |
lemma (in prob_space) indep_setsI: |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
85 |
assumes "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
86 |
and "\<And>A J. J \<noteq> {} \<Longrightarrow> J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> (\<forall>j\<in>J. A j \<in> F j) \<Longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
87 |
shows "indep_sets F I" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
88 |
using assms unfolding indep_sets_def by (auto simp: Pi_iff) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
89 |
|
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
90 |
lemma (in prob_space) indep_setsD: |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
91 |
assumes "indep_sets F I" and "J \<subseteq> I" "J \<noteq> {}" "finite J" "\<forall>j\<in>J. A j \<in> F j" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
92 |
shows "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
93 |
using assms unfolding indep_sets_def by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
94 |
|
42982 | 95 |
lemma (in prob_space) indep_setI: |
96 |
assumes ev: "A \<subseteq> events" "B \<subseteq> events" |
|
97 |
and indep: "\<And>a b. a \<in> A \<Longrightarrow> b \<in> B \<Longrightarrow> prob (a \<inter> b) = prob a * prob b" |
|
98 |
shows "indep_set A B" |
|
99 |
unfolding indep_set_def |
|
100 |
proof (rule indep_setsI) |
|
101 |
fix F J assume "J \<noteq> {}" "J \<subseteq> UNIV" |
|
102 |
and F: "\<forall>j\<in>J. F j \<in> (case j of True \<Rightarrow> A | False \<Rightarrow> B)" |
|
103 |
have "J \<in> Pow UNIV" by auto |
|
61808 | 104 |
with F \<open>J \<noteq> {}\<close> indep[of "F True" "F False"] |
42982 | 105 |
show "prob (\<Inter>j\<in>J. F j) = (\<Prod>j\<in>J. prob (F j))" |
106 |
unfolding UNIV_bool Pow_insert by (auto simp: ac_simps) |
|
107 |
qed (auto split: bool.split simp: ev) |
|
108 |
||
109 |
lemma (in prob_space) indep_setD: |
|
110 |
assumes indep: "indep_set A B" and ev: "a \<in> A" "b \<in> B" |
|
111 |
shows "prob (a \<inter> b) = prob a * prob b" |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
112 |
using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "case_bool a b"] ev |
42982 | 113 |
by (simp add: ac_simps UNIV_bool) |
114 |
||
115 |
lemma (in prob_space) |
|
116 |
assumes indep: "indep_set A B" |
|
42983 | 117 |
shows indep_setD_ev1: "A \<subseteq> events" |
118 |
and indep_setD_ev2: "B \<subseteq> events" |
|
42982 | 119 |
using indep unfolding indep_set_def indep_sets_def UNIV_bool by auto |
120 |
||
69555 | 121 |
lemma (in prob_space) indep_sets_Dynkin: |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
122 |
assumes indep: "indep_sets F I" |
69555 | 123 |
shows "indep_sets (\<lambda>i. Dynkin (space M) (F i)) I" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
124 |
(is "indep_sets ?F I") |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
125 |
proof (subst indep_sets_finite_index_sets, intro allI impI ballI) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
126 |
fix J assume "finite J" "J \<subseteq> I" "J \<noteq> {}" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
127 |
with indep have "indep_sets F J" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
128 |
by (subst (asm) indep_sets_finite_index_sets) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
129 |
{ fix J K assume "indep_sets F K" |
46731 | 130 |
let ?G = "\<lambda>S i. if i \<in> S then ?F i else F i" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
131 |
assume "finite J" "J \<subseteq> K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
132 |
then have "indep_sets (?G J) K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
133 |
proof induct |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
134 |
case (insert j J) |
63040 | 135 |
moreover define G where "G = ?G J" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
136 |
ultimately have G: "indep_sets G K" "\<And>i. i \<in> K \<Longrightarrow> G i \<subseteq> events" and "j \<in> K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
137 |
by (auto simp: indep_sets_def) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
138 |
let ?D = "{E\<in>events. indep_sets (G(j := {E})) K }" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
139 |
{ fix X assume X: "X \<in> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
140 |
assume indep: "\<And>J A. J \<noteq> {} \<Longrightarrow> J \<subseteq> K \<Longrightarrow> finite J \<Longrightarrow> j \<notin> J \<Longrightarrow> (\<forall>i\<in>J. A i \<in> G i) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
141 |
\<Longrightarrow> prob ((\<Inter>i\<in>J. A i) \<inter> X) = prob X * (\<Prod>i\<in>J. prob (A i))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
142 |
have "indep_sets (G(j := {X})) K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
143 |
proof (rule indep_setsI) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
144 |
fix i assume "i \<in> K" then show "(G(j:={X})) i \<subseteq> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
145 |
using G X by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
146 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
147 |
fix A J assume J: "J \<noteq> {}" "J \<subseteq> K" "finite J" "\<forall>i\<in>J. A i \<in> (G(j := {X})) i" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
148 |
show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
149 |
proof cases |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
150 |
assume "j \<in> J" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
151 |
with J have "A j = X" by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
152 |
show ?thesis |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
153 |
proof cases |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
154 |
assume "J = {j}" then show ?thesis by simp |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
155 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
156 |
assume "J \<noteq> {j}" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
157 |
have "prob (\<Inter>i\<in>J. A i) = prob ((\<Inter>i\<in>J-{j}. A i) \<inter> X)" |
62390 | 158 |
using \<open>j \<in> J\<close> \<open>A j = X\<close> by (auto intro!: arg_cong[where f=prob] split: if_split_asm) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
159 |
also have "\<dots> = prob X * (\<Prod>i\<in>J-{j}. prob (A i))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
160 |
proof (rule indep) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
161 |
show "J - {j} \<noteq> {}" "J - {j} \<subseteq> K" "finite (J - {j})" "j \<notin> J - {j}" |
61808 | 162 |
using J \<open>J \<noteq> {j}\<close> \<open>j \<in> J\<close> by auto |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
163 |
show "\<forall>i\<in>J - {j}. A i \<in> G i" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
164 |
using J by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
165 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
166 |
also have "\<dots> = prob (A j) * (\<Prod>i\<in>J-{j}. prob (A i))" |
61808 | 167 |
using \<open>A j = X\<close> by simp |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
168 |
also have "\<dots> = (\<Prod>i\<in>J. prob (A i))" |
64272 | 169 |
unfolding prod.insert_remove[OF \<open>finite J\<close>, symmetric, of "\<lambda>i. prob (A i)"] |
61808 | 170 |
using \<open>j \<in> J\<close> by (simp add: insert_absorb) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
171 |
finally show ?thesis . |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
172 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
173 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
174 |
assume "j \<notin> J" |
62390 | 175 |
with J have "\<forall>i\<in>J. A i \<in> G i" by (auto split: if_split_asm) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
176 |
with J show ?thesis |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
177 |
by (intro indep_setsD[OF G(1)]) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
178 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
179 |
qed } |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
180 |
note indep_sets_insert = this |
69555 | 181 |
have "Dynkin_system (space M) ?D" |
182 |
proof (rule Dynkin_systemI', simp_all cong del: indep_sets_cong, safe) |
|
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
183 |
show "indep_sets (G(j := {{}})) K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
184 |
by (rule indep_sets_insert) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
185 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
186 |
fix X assume X: "X \<in> events" and G': "indep_sets (G(j := {X})) K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
187 |
show "indep_sets (G(j := {space M - X})) K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
188 |
proof (rule indep_sets_insert) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
189 |
fix J A assume J: "J \<noteq> {}" "J \<subseteq> K" "finite J" "j \<notin> J" and A: "\<forall>i\<in>J. A i \<in> G i" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
190 |
then have A_sets: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
191 |
using G by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
192 |
have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
193 |
prob ((\<Inter>j\<in>J. A j) - (\<Inter>i\<in>insert j J. (A(j := X)) i))" |
61808 | 194 |
using A_sets sets.sets_into_space[of _ M] X \<open>J \<noteq> {}\<close> |
62390 | 195 |
by (auto intro!: arg_cong[where f=prob] split: if_split_asm) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
196 |
also have "\<dots> = prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)" |
61808 | 197 |
using J \<open>J \<noteq> {}\<close> \<open>j \<notin> J\<close> A_sets X sets.sets_into_space |
62390 | 198 |
by (auto intro!: finite_measure_Diff sets.finite_INT split: if_split_asm) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
199 |
finally have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
200 |
prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)" . |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
201 |
moreover { |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
202 |
have "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" |
61808 | 203 |
using J A \<open>finite J\<close> by (intro indep_setsD[OF G(1)]) auto |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
204 |
then have "prob (\<Inter>j\<in>J. A j) = prob (space M) * (\<Prod>i\<in>J. prob (A i))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
205 |
using prob_space by simp } |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
206 |
moreover { |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
207 |
have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = (\<Prod>i\<in>insert j J. prob ((A(j := X)) i))" |
61808 | 208 |
using J A \<open>j \<in> K\<close> by (intro indep_setsD[OF G']) auto |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
209 |
then have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = prob X * (\<Prod>i\<in>J. prob (A i))" |
64272 | 210 |
using \<open>finite J\<close> \<open>j \<notin> J\<close> by (auto intro!: prod.cong) } |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
211 |
ultimately have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = (prob (space M) - prob X) * (\<Prod>i\<in>J. prob (A i))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
212 |
by (simp add: field_simps) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
213 |
also have "\<dots> = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
214 |
using X A by (simp add: finite_measure_compl) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
215 |
finally show "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))" . |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
216 |
qed (insert X, auto) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
217 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
218 |
fix F :: "nat \<Rightarrow> 'a set" assume disj: "disjoint_family F" and "range F \<subseteq> ?D" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
219 |
then have F: "\<And>i. F i \<in> events" "\<And>i. indep_sets (G(j:={F i})) K" by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
220 |
show "indep_sets (G(j := {\<Union>k. F k})) K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
221 |
proof (rule indep_sets_insert) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
222 |
fix J A assume J: "j \<notin> J" "J \<noteq> {}" "J \<subseteq> K" "finite J" and A: "\<forall>i\<in>J. A i \<in> G i" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
223 |
then have A_sets: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
224 |
using G by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
225 |
have "prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)) = prob (\<Union>k. (\<Inter>i\<in>insert j J. (A(j := F k)) i))" |
62390 | 226 |
using \<open>J \<noteq> {}\<close> \<open>j \<notin> J\<close> \<open>j \<in> K\<close> by (auto intro!: arg_cong[where f=prob] split: if_split_asm) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
227 |
moreover have "(\<lambda>k. prob (\<Inter>i\<in>insert j J. (A(j := F k)) i)) sums prob (\<Union>k. (\<Inter>i\<in>insert j J. (A(j := F k)) i))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
228 |
proof (rule finite_measure_UNION) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
229 |
show "disjoint_family (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i)" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
230 |
using disj by (rule disjoint_family_on_bisimulation) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
231 |
show "range (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i) \<subseteq> events" |
61808 | 232 |
using A_sets F \<open>finite J\<close> \<open>J \<noteq> {}\<close> \<open>j \<notin> J\<close> by (auto intro!: sets.Int) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
233 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
234 |
moreover { fix k |
61808 | 235 |
from J A \<open>j \<in> K\<close> have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * (\<Prod>i\<in>J. prob (A i))" |
64272 | 236 |
by (subst indep_setsD[OF F(2)]) (auto intro!: prod.cong split: if_split_asm) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
237 |
also have "\<dots> = prob (F k) * prob (\<Inter>i\<in>J. A i)" |
61808 | 238 |
using J A \<open>j \<in> K\<close> by (subst indep_setsD[OF G(1)]) auto |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
239 |
finally have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * prob (\<Inter>i\<in>J. A i)" . } |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
240 |
ultimately have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
241 |
by simp |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
242 |
moreover |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
243 |
have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob (\<Union>k. F k) * prob (\<Inter>i\<in>J. A i))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
244 |
using disj F(1) by (intro finite_measure_UNION sums_mult2) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
245 |
then have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob (\<Union>k. F k) * (\<Prod>i\<in>J. prob (A i)))" |
61808 | 246 |
using J A \<open>j \<in> K\<close> by (subst indep_setsD[OF G(1), symmetric]) auto |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
247 |
ultimately |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
248 |
show "prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)) = prob (\<Union>k. F k) * (\<Prod>j\<in>J. prob (A j))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
249 |
by (auto dest!: sums_unique) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
250 |
qed (insert F, auto) |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
251 |
qed (insert sets.sets_into_space, auto) |
69555 | 252 |
then have mono: "Dynkin (space M) (G j) \<subseteq> {E \<in> events. indep_sets (G(j := {E})) K}" |
253 |
proof (rule Dynkin_system.Dynkin_subset, safe) |
|
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
254 |
fix X assume "X \<in> G j" |
61808 | 255 |
then show "X \<in> events" using G \<open>j \<in> K\<close> by auto |
256 |
from \<open>indep_sets G K\<close> |
|
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
257 |
show "indep_sets (G(j := {X})) K" |
61808 | 258 |
by (rule indep_sets_mono_sets) (insert \<open>X \<in> G j\<close>, auto) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
259 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
260 |
have "indep_sets (G(j:=?D)) K" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
261 |
proof (rule indep_setsI) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
262 |
fix i assume "i \<in> K" then show "(G(j := ?D)) i \<subseteq> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
263 |
using G(2) by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
264 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
265 |
fix A J assume J: "J\<noteq>{}" "J \<subseteq> K" "finite J" and A: "\<forall>i\<in>J. A i \<in> (G(j := ?D)) i" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
266 |
show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
267 |
proof cases |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
268 |
assume "j \<in> J" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
269 |
with A have indep: "indep_sets (G(j := {A j})) K" by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
270 |
from J A show ?thesis |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
271 |
by (intro indep_setsD[OF indep]) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
272 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
273 |
assume "j \<notin> J" |
62390 | 274 |
with J A have "\<forall>i\<in>J. A i \<in> G i" by (auto split: if_split_asm) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
275 |
with J show ?thesis |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
276 |
by (intro indep_setsD[OF G(1)]) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
277 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
278 |
qed |
69555 | 279 |
then have "indep_sets (G(j := Dynkin (space M) (G j))) K" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
280 |
by (rule indep_sets_mono_sets) (insert mono, auto) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
281 |
then show ?case |
61808 | 282 |
by (rule indep_sets_mono_sets) (insert \<open>j \<in> K\<close> \<open>j \<notin> J\<close>, auto simp: G_def) |
283 |
qed (insert \<open>indep_sets F K\<close>, simp) } |
|
284 |
from this[OF \<open>indep_sets F J\<close> \<open>finite J\<close> subset_refl] |
|
47694 | 285 |
show "indep_sets ?F J" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
286 |
by (rule indep_sets_mono_sets) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
287 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
288 |
|
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
289 |
lemma (in prob_space) indep_sets_sigma: |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
290 |
assumes indep: "indep_sets F I" |
47694 | 291 |
assumes stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (F i)" |
292 |
shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I" |
|
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
293 |
proof - |
69555 | 294 |
from indep_sets_Dynkin[OF indep] |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
295 |
show ?thesis |
69555 | 296 |
proof (rule indep_sets_mono_sets, subst sigma_eq_Dynkin, simp_all add: stable) |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
297 |
fix i assume "i \<in> I" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
298 |
with indep have "F i \<subseteq> events" by (auto simp: indep_sets_def) |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
299 |
with sets.sets_into_space show "F i \<subseteq> Pow (space M)" by auto |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
300 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
301 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
302 |
|
42987 | 303 |
lemma (in prob_space) indep_sets_sigma_sets_iff: |
47694 | 304 |
assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable (F i)" |
42987 | 305 |
shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I \<longleftrightarrow> indep_sets F I" |
306 |
proof |
|
307 |
assume "indep_sets F I" then show "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I" |
|
47694 | 308 |
by (rule indep_sets_sigma) fact |
42987 | 309 |
next |
310 |
assume "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I" then show "indep_sets F I" |
|
311 |
by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic) |
|
312 |
qed |
|
313 |
||
49794 | 314 |
definition (in prob_space) |
315 |
indep_vars_def2: "indep_vars M' X I \<longleftrightarrow> |
|
49781 | 316 |
(\<forall>i\<in>I. random_variable (M' i) (X i)) \<and> |
317 |
indep_sets (\<lambda>i. { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I" |
|
49794 | 318 |
|
319 |
definition (in prob_space) |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
320 |
"indep_var Ma A Mb B \<longleftrightarrow> indep_vars (case_bool Ma Mb) (case_bool A B) UNIV" |
49794 | 321 |
|
322 |
lemma (in prob_space) indep_vars_def: |
|
323 |
"indep_vars M' X I \<longleftrightarrow> |
|
324 |
(\<forall>i\<in>I. random_variable (M' i) (X i)) \<and> |
|
325 |
indep_sets (\<lambda>i. sigma_sets (space M) { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I" |
|
326 |
unfolding indep_vars_def2 |
|
49781 | 327 |
apply (rule conj_cong[OF refl]) |
49794 | 328 |
apply (rule indep_sets_sigma_sets_iff[symmetric]) |
49781 | 329 |
apply (auto simp: Int_stable_def) |
330 |
apply (rule_tac x="A \<inter> Aa" in exI) |
|
331 |
apply auto |
|
332 |
done |
|
333 |
||
49794 | 334 |
lemma (in prob_space) indep_var_eq: |
335 |
"indep_var S X T Y \<longleftrightarrow> |
|
336 |
(random_variable S X \<and> random_variable T Y) \<and> |
|
337 |
indep_set |
|
338 |
(sigma_sets (space M) { X -` A \<inter> space M | A. A \<in> sets S}) |
|
339 |
(sigma_sets (space M) { Y -` A \<inter> space M | A. A \<in> sets T})" |
|
340 |
unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool |
|
67399 | 341 |
by (intro arg_cong2[where f="(\<and>)"] arg_cong2[where f=indep_sets] ext) |
49794 | 342 |
(auto split: bool.split) |
343 |
||
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
344 |
lemma (in prob_space) indep_sets2_eq: |
42981 | 345 |
"indep_set A B \<longleftrightarrow> A \<subseteq> events \<and> B \<subseteq> events \<and> (\<forall>a\<in>A. \<forall>b\<in>B. prob (a \<inter> b) = prob a * prob b)" |
346 |
unfolding indep_set_def |
|
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
347 |
proof (intro iffI ballI conjI) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
348 |
assume indep: "indep_sets (case_bool A B) UNIV" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
349 |
{ fix a b assume "a \<in> A" "b \<in> B" |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
350 |
with indep_setsD[OF indep, of UNIV "case_bool a b"] |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
351 |
show "prob (a \<inter> b) = prob a * prob b" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
352 |
unfolding UNIV_bool by (simp add: ac_simps) } |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
353 |
from indep show "A \<subseteq> events" "B \<subseteq> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
354 |
unfolding indep_sets_def UNIV_bool by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
355 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
356 |
assume *: "A \<subseteq> events \<and> B \<subseteq> events \<and> (\<forall>a\<in>A. \<forall>b\<in>B. prob (a \<inter> b) = prob a * prob b)" |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
357 |
show "indep_sets (case_bool A B) UNIV" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
358 |
proof (rule indep_setsI) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
359 |
fix i show "(case i of True \<Rightarrow> A | False \<Rightarrow> B) \<subseteq> events" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
360 |
using * by (auto split: bool.split) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
361 |
next |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
362 |
fix J X assume "J \<noteq> {}" "J \<subseteq> UNIV" and X: "\<forall>j\<in>J. X j \<in> (case j of True \<Rightarrow> A | False \<Rightarrow> B)" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
363 |
then have "J = {True} \<or> J = {False} \<or> J = {True,False}" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
364 |
by (auto simp: UNIV_bool) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
365 |
then show "prob (\<Inter>j\<in>J. X j) = (\<Prod>j\<in>J. prob (X j))" |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
366 |
using X * by auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
367 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
368 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
369 |
|
42981 | 370 |
lemma (in prob_space) indep_set_sigma_sets: |
371 |
assumes "indep_set A B" |
|
47694 | 372 |
assumes A: "Int_stable A" and B: "Int_stable B" |
42981 | 373 |
shows "indep_set (sigma_sets (space M) A) (sigma_sets (space M) B)" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
374 |
proof - |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
375 |
have "indep_sets (\<lambda>i. sigma_sets (space M) (case i of True \<Rightarrow> A | False \<Rightarrow> B)) UNIV" |
47694 | 376 |
proof (rule indep_sets_sigma) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
377 |
show "indep_sets (case_bool A B) UNIV" |
61808 | 378 |
by (rule \<open>indep_set A B\<close>[unfolded indep_set_def]) |
47694 | 379 |
fix i show "Int_stable (case i of True \<Rightarrow> A | False \<Rightarrow> B)" |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
380 |
using A B by (cases i) auto |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
381 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
382 |
then show ?thesis |
42981 | 383 |
unfolding indep_set_def |
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
384 |
by (rule indep_sets_mono_sets) (auto split: bool.split) |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
385 |
qed |
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
386 |
|
59000 | 387 |
lemma (in prob_space) indep_eventsI_indep_vars: |
388 |
assumes indep: "indep_vars N X I" |
|
389 |
assumes P: "\<And>i. i \<in> I \<Longrightarrow> {x\<in>space (N i). P i x} \<in> sets (N i)" |
|
390 |
shows "indep_events (\<lambda>i. {x\<in>space M. P i (X i x)}) I" |
|
391 |
proof - |
|
392 |
have "indep_sets (\<lambda>i. {X i -` A \<inter> space M |A. A \<in> sets (N i)}) I" |
|
393 |
using indep unfolding indep_vars_def2 by auto |
|
394 |
then show ?thesis |
|
395 |
unfolding indep_events_def_alt |
|
396 |
proof (rule indep_sets_mono_sets) |
|
397 |
fix i assume "i \<in> I" |
|
398 |
then have "{{x \<in> space M. P i (X i x)}} = {X i -` {x\<in>space (N i). P i x} \<inter> space M}" |
|
399 |
using indep by (auto simp: indep_vars_def dest: measurable_space) |
|
400 |
also have "\<dots> \<subseteq> {X i -` A \<inter> space M |A. A \<in> sets (N i)}" |
|
61808 | 401 |
using P[OF \<open>i \<in> I\<close>] by blast |
59000 | 402 |
finally show "{{x \<in> space M. P i (X i x)}} \<subseteq> {X i -` A \<inter> space M |A. A \<in> sets (N i)}" . |
403 |
qed |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
404 |
qed |
59000 | 405 |
|
42981 | 406 |
lemma (in prob_space) indep_sets_collect_sigma: |
407 |
fixes I :: "'j \<Rightarrow> 'i set" and J :: "'j set" and E :: "'i \<Rightarrow> 'a set set" |
|
408 |
assumes indep: "indep_sets E (\<Union>j\<in>J. I j)" |
|
47694 | 409 |
assumes Int_stable: "\<And>i j. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> Int_stable (E i)" |
42981 | 410 |
assumes disjoint: "disjoint_family_on I J" |
411 |
shows "indep_sets (\<lambda>j. sigma_sets (space M) (\<Union>i\<in>I j. E i)) J" |
|
412 |
proof - |
|
46731 | 413 |
let ?E = "\<lambda>j. {\<Inter>k\<in>K. E' k| E' K. finite K \<and> K \<noteq> {} \<and> K \<subseteq> I j \<and> (\<forall>k\<in>K. E' k \<in> E k) }" |
42981 | 414 |
|
42983 | 415 |
from indep have E: "\<And>j i. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> E i \<subseteq> events" |
42981 | 416 |
unfolding indep_sets_def by auto |
417 |
{ fix j |
|
47694 | 418 |
let ?S = "sigma_sets (space M) (\<Union>i\<in>I j. E i)" |
42981 | 419 |
assume "j \<in> J" |
47694 | 420 |
from E[OF this] interpret S: sigma_algebra "space M" ?S |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
421 |
using sets.sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto |
42981 | 422 |
|
423 |
have "sigma_sets (space M) (\<Union>i\<in>I j. E i) = sigma_sets (space M) (?E j)" |
|
424 |
proof (rule sigma_sets_eqI) |
|
425 |
fix A assume "A \<in> (\<Union>i\<in>I j. E i)" |
|
74362 | 426 |
then obtain i where "i \<in> I j" "A \<in> E i" .. |
42981 | 427 |
then show "A \<in> sigma_sets (space M) (?E j)" |
47694 | 428 |
by (auto intro!: sigma_sets.intros(2-) exI[of _ "{i}"] exI[of _ "\<lambda>i. A"]) |
42981 | 429 |
next |
430 |
fix A assume "A \<in> ?E j" |
|
431 |
then obtain E' K where "finite K" "K \<noteq> {}" "K \<subseteq> I j" "\<And>k. k \<in> K \<Longrightarrow> E' k \<in> E k" |
|
432 |
and A: "A = (\<Inter>k\<in>K. E' k)" |
|
433 |
by auto |
|
47694 | 434 |
then have "A \<in> ?S" unfolding A |
435 |
by (safe intro!: S.finite_INT) auto |
|
42981 | 436 |
then show "A \<in> sigma_sets (space M) (\<Union>i\<in>I j. E i)" |
47694 | 437 |
by simp |
42981 | 438 |
qed } |
439 |
moreover have "indep_sets (\<lambda>j. sigma_sets (space M) (?E j)) J" |
|
47694 | 440 |
proof (rule indep_sets_sigma) |
42981 | 441 |
show "indep_sets ?E J" |
442 |
proof (intro indep_setsI) |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
443 |
fix j assume "j \<in> J" with E show "?E j \<subseteq> events" by (force intro!: sets.finite_INT) |
42981 | 444 |
next |
445 |
fix K A assume K: "K \<noteq> {}" "K \<subseteq> J" "finite K" |
|
446 |
and "\<forall>j\<in>K. A j \<in> ?E j" |
|
447 |
then have "\<forall>j\<in>K. \<exists>E' L. A j = (\<Inter>l\<in>L. E' l) \<and> finite L \<and> L \<noteq> {} \<and> L \<subseteq> I j \<and> (\<forall>l\<in>L. E' l \<in> E l)" |
|
448 |
by simp |
|
74362 | 449 |
from bchoice[OF this] obtain E' |
450 |
where "\<forall>x\<in>K. \<exists>L. A x = \<Inter> (E' x ` L) \<and> finite L \<and> L \<noteq> {} \<and> L \<subseteq> I x \<and> (\<forall>l\<in>L. E' x l \<in> E l)" |
|
451 |
.. |
|
42981 | 452 |
from bchoice[OF this] obtain L |
453 |
where A: "\<And>j. j\<in>K \<Longrightarrow> A j = (\<Inter>l\<in>L j. E' j l)" |
|
454 |
and L: "\<And>j. j\<in>K \<Longrightarrow> finite (L j)" "\<And>j. j\<in>K \<Longrightarrow> L j \<noteq> {}" "\<And>j. j\<in>K \<Longrightarrow> L j \<subseteq> I j" |
|
455 |
and E': "\<And>j l. j\<in>K \<Longrightarrow> l \<in> L j \<Longrightarrow> E' j l \<in> E l" |
|
456 |
by auto |
|
457 |
{ fix k l j assume "k \<in> K" "j \<in> K" "l \<in> L j" "l \<in> L k" |
|
458 |
have "k = j" |
|
459 |
proof (rule ccontr) |
|
460 |
assume "k \<noteq> j" |
|
61808 | 461 |
with disjoint \<open>K \<subseteq> J\<close> \<open>k \<in> K\<close> \<open>j \<in> K\<close> have "I k \<inter> I j = {}" |
42981 | 462 |
unfolding disjoint_family_on_def by auto |
61808 | 463 |
with L(2,3)[OF \<open>j \<in> K\<close>] L(2,3)[OF \<open>k \<in> K\<close>] |
464 |
show False using \<open>l \<in> L k\<close> \<open>l \<in> L j\<close> by auto |
|
42981 | 465 |
qed } |
466 |
note L_inj = this |
|
467 |
||
63040 | 468 |
define k where "k l = (SOME k. k \<in> K \<and> l \<in> L k)" for l |
42981 | 469 |
{ fix x j l assume *: "j \<in> K" "l \<in> L j" |
470 |
have "k l = j" unfolding k_def |
|
471 |
proof (rule some_equality) |
|
472 |
fix k assume "k \<in> K \<and> l \<in> L k" |
|
473 |
with * L_inj show "k = j" by auto |
|
474 |
qed (insert *, simp) } |
|
475 |
note k_simp[simp] = this |
|
46731 | 476 |
let ?E' = "\<lambda>l. E' (k l) l" |
42981 | 477 |
have "prob (\<Inter>j\<in>K. A j) = prob (\<Inter>l\<in>(\<Union>k\<in>K. L k). ?E' l)" |
478 |
by (auto simp: A intro!: arg_cong[where f=prob]) |
|
479 |
also have "\<dots> = (\<Prod>l\<in>(\<Union>k\<in>K. L k). prob (?E' l))" |
|
480 |
using L K E' by (intro indep_setsD[OF indep]) (simp_all add: UN_mono) |
|
481 |
also have "\<dots> = (\<Prod>j\<in>K. \<Prod>l\<in>L j. prob (E' j l))" |
|
64272 | 482 |
using K L L_inj by (subst prod.UNION_disjoint) auto |
42981 | 483 |
also have "\<dots> = (\<Prod>j\<in>K. prob (A j))" |
64272 | 484 |
using K L E' by (auto simp add: A intro!: prod.cong indep_setsD[OF indep, symmetric]) blast |
42981 | 485 |
finally show "prob (\<Inter>j\<in>K. A j) = (\<Prod>j\<in>K. prob (A j))" . |
486 |
qed |
|
487 |
next |
|
488 |
fix j assume "j \<in> J" |
|
47694 | 489 |
show "Int_stable (?E j)" |
42981 | 490 |
proof (rule Int_stableI) |
491 |
fix a assume "a \<in> ?E j" then obtain Ka Ea |
|
492 |
where a: "a = (\<Inter>k\<in>Ka. Ea k)" "finite Ka" "Ka \<noteq> {}" "Ka \<subseteq> I j" "\<And>k. k\<in>Ka \<Longrightarrow> Ea k \<in> E k" by auto |
|
493 |
fix b assume "b \<in> ?E j" then obtain Kb Eb |
|
494 |
where b: "b = (\<Inter>k\<in>Kb. Eb k)" "finite Kb" "Kb \<noteq> {}" "Kb \<subseteq> I j" "\<And>k. k\<in>Kb \<Longrightarrow> Eb k \<in> E k" by auto |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
495 |
let ?f = "\<lambda>k. (if k \<in> Ka \<inter> Kb then Ea k \<inter> Eb k else if k \<in> Kb then Eb k else if k \<in> Ka then Ea k else {})" |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
496 |
have "Ka \<union> Kb = (Ka \<inter> Kb) \<union> (Kb - Ka) \<union> (Ka - Kb)" |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
497 |
by blast |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
498 |
moreover have "(\<Inter>x\<in>Ka \<inter> Kb. Ea x \<inter> Eb x) \<inter> |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
499 |
(\<Inter>x\<in>Kb - Ka. Eb x) \<inter> (\<Inter>x\<in>Ka - Kb. Ea x) = (\<Inter>k\<in>Ka. Ea k) \<inter> (\<Inter>k\<in>Kb. Eb k)" |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
500 |
by auto |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
501 |
ultimately have "(\<Inter>k\<in>Ka \<union> Kb. ?f k) = (\<Inter>k\<in>Ka. Ea k) \<inter> (\<Inter>k\<in>Kb. Eb k)" (is "?lhs = ?rhs") |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
502 |
by (simp only: image_Un Inter_Un_distrib) simp |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
503 |
then have "a \<inter> b = (\<Inter>k\<in>Ka \<union> Kb. ?f k)" |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
504 |
by (simp only: a(1) b(1)) |
61808 | 505 |
with a b \<open>j \<in> J\<close> Int_stableD[OF Int_stable] show "a \<inter> b \<in> ?E j" |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
506 |
by (intro CollectI exI[of _ "Ka \<union> Kb"] exI[of _ ?f]) auto |
42981 | 507 |
qed |
508 |
qed |
|
509 |
ultimately show ?thesis |
|
510 |
by (simp cong: indep_sets_cong) |
|
511 |
qed |
|
512 |
||
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
513 |
lemma (in prob_space) indep_vars_restrict: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
514 |
assumes ind: "indep_vars M' X I" and K: "\<And>j. j \<in> L \<Longrightarrow> K j \<subseteq> I" and J: "disjoint_family_on K L" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
515 |
shows "indep_vars (\<lambda>j. PiM (K j) M') (\<lambda>j \<omega>. restrict (\<lambda>i. X i \<omega>) (K j)) L" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
516 |
unfolding indep_vars_def |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
517 |
proof safe |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
518 |
fix j assume "j \<in> L" then show "random_variable (Pi\<^sub>M (K j) M') (\<lambda>\<omega>. \<lambda>i\<in>K j. X i \<omega>)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
519 |
using K ind by (auto simp: indep_vars_def intro!: measurable_restrict) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
520 |
next |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
521 |
have X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable M (M' i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
522 |
using ind by (auto simp: indep_vars_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
523 |
let ?proj = "\<lambda>j S. {(\<lambda>\<omega>. \<lambda>i\<in>K j. X i \<omega>) -` A \<inter> space M |A. A \<in> S}" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
524 |
let ?UN = "\<lambda>j. sigma_sets (space M) (\<Union>i\<in>K j. { X i -` A \<inter> space M| A. A \<in> sets (M' i) })" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
525 |
show "indep_sets (\<lambda>i. sigma_sets (space M) (?proj i (sets (Pi\<^sub>M (K i) M')))) L" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
526 |
proof (rule indep_sets_mono_sets) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
527 |
fix j assume j: "j \<in> L" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
528 |
have "sigma_sets (space M) (?proj j (sets (Pi\<^sub>M (K j) M'))) = |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
529 |
sigma_sets (space M) (sigma_sets (space M) (?proj j (prod_algebra (K j) M')))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
530 |
using j K X[THEN measurable_space] unfolding sets_PiM |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
531 |
by (subst sigma_sets_vimage_commute) (auto simp add: Pi_iff) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
532 |
also have "\<dots> = sigma_sets (space M) (?proj j (prod_algebra (K j) M'))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
533 |
by (rule sigma_sets_sigma_sets_eq) auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
534 |
also have "\<dots> \<subseteq> ?UN j" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
535 |
proof (rule sigma_sets_mono, safe del: disjE elim!: prod_algebraE) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
536 |
fix J E assume J: "finite J" "J \<noteq> {} \<or> K j = {}" "J \<subseteq> K j" and E: "\<forall>i. i \<in> J \<longrightarrow> E i \<in> sets (M' i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
537 |
show "(\<lambda>\<omega>. \<lambda>i\<in>K j. X i \<omega>) -` prod_emb (K j) M' J (Pi\<^sub>E J E) \<inter> space M \<in> ?UN j" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
538 |
proof cases |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
539 |
assume "K j = {}" with J show ?thesis |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
540 |
by (auto simp add: sigma_sets_empty_eq prod_emb_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
541 |
next |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
542 |
assume "K j \<noteq> {}" with J have "J \<noteq> {}" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
543 |
by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
544 |
{ interpret sigma_algebra "space M" "?UN j" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
545 |
by (rule sigma_algebra_sigma_sets) auto |
69313 | 546 |
have "\<And>A. (\<And>i. i \<in> J \<Longrightarrow> A i \<in> ?UN j) \<Longrightarrow> \<Inter>(A ` J) \<in> ?UN j" |
61808 | 547 |
using \<open>finite J\<close> \<open>J \<noteq> {}\<close> by (rule finite_INT) blast } |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
548 |
note INT = this |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
549 |
|
61808 | 550 |
from \<open>J \<noteq> {}\<close> J K E[rule_format, THEN sets.sets_into_space] j |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
551 |
have "(\<lambda>\<omega>. \<lambda>i\<in>K j. X i \<omega>) -` prod_emb (K j) M' J (Pi\<^sub>E J E) \<inter> space M |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
552 |
= (\<Inter>i\<in>J. X i -` E i \<inter> space M)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
553 |
apply (subst prod_emb_PiE[OF _ ]) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
554 |
apply auto [] |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
555 |
apply auto [] |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
556 |
apply (auto simp add: Pi_iff intro!: X[THEN measurable_space]) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
557 |
apply (erule_tac x=i in ballE) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
558 |
apply auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
559 |
done |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
560 |
also have "\<dots> \<in> ?UN j" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
561 |
apply (rule INT) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
562 |
apply (rule sigma_sets.Basic) |
61808 | 563 |
using \<open>J \<subseteq> K j\<close> E |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
564 |
apply auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
565 |
done |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
566 |
finally show ?thesis . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
567 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
568 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
569 |
finally show "sigma_sets (space M) (?proj j (sets (Pi\<^sub>M (K j) M'))) \<subseteq> ?UN j" . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
570 |
next |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
571 |
show "indep_sets ?UN L" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
572 |
proof (rule indep_sets_collect_sigma) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
573 |
show "indep_sets (\<lambda>i. {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) (\<Union>j\<in>L. K j)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
574 |
proof (rule indep_sets_mono_index) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
575 |
show "indep_sets (\<lambda>i. {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
576 |
using ind unfolding indep_vars_def2 by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
577 |
show "(\<Union>l\<in>L. K l) \<subseteq> I" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
578 |
using K by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
579 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
580 |
next |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
581 |
fix l i assume "l \<in> L" "i \<in> K l" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
582 |
show "Int_stable {X i -` A \<inter> space M |A. A \<in> sets (M' i)}" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
583 |
apply (auto simp: Int_stable_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
584 |
apply (rule_tac x="A \<inter> Aa" in exI) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
585 |
apply auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
586 |
done |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
587 |
qed fact |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
588 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
589 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
590 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
591 |
lemma (in prob_space) indep_var_restrict: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
592 |
assumes ind: "indep_vars M' X I" and AB: "A \<inter> B = {}" "A \<subseteq> I" "B \<subseteq> I" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
593 |
shows "indep_var (PiM A M') (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) A) (PiM B M') (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) B)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
594 |
proof - |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
595 |
have *: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
596 |
"case_bool (Pi\<^sub>M A M') (Pi\<^sub>M B M') = (\<lambda>b. PiM (case_bool A B b) M')" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
597 |
"case_bool (\<lambda>\<omega>. \<lambda>i\<in>A. X i \<omega>) (\<lambda>\<omega>. \<lambda>i\<in>B. X i \<omega>) = (\<lambda>b \<omega>. \<lambda>i\<in>case_bool A B b. X i \<omega>)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
598 |
by (simp_all add: fun_eq_iff split: bool.split) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
599 |
show ?thesis |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
600 |
unfolding indep_var_def * using AB |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
601 |
by (intro indep_vars_restrict[OF ind]) (auto simp: disjoint_family_on_def split: bool.split) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
602 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
603 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
604 |
lemma (in prob_space) indep_vars_subset: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
605 |
assumes "indep_vars M' X I" "J \<subseteq> I" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
606 |
shows "indep_vars M' X J" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
607 |
using assms unfolding indep_vars_def indep_sets_def |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
608 |
by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
609 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
610 |
lemma (in prob_space) indep_vars_cong: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
611 |
"I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> X i = Y i) \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> M' i = N' i) \<Longrightarrow> indep_vars M' X I \<longleftrightarrow> indep_vars N' Y J" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
612 |
unfolding indep_vars_def2 by (intro conj_cong indep_sets_cong) auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
613 |
|
49772 | 614 |
definition (in prob_space) tail_events where |
69325 | 615 |
"tail_events A = (\<Inter>n. sigma_sets (space M) (\<Union> (A ` {n..})))" |
42982 | 616 |
|
49772 | 617 |
lemma (in prob_space) tail_events_sets: |
618 |
assumes A: "\<And>i::nat. A i \<subseteq> events" |
|
619 |
shows "tail_events A \<subseteq> events" |
|
620 |
proof |
|
621 |
fix X assume X: "X \<in> tail_events A" |
|
69325 | 622 |
let ?A = "(\<Inter>n. sigma_sets (space M) (\<Union> (A ` {n..})))" |
623 |
from X have "\<And>n::nat. X \<in> sigma_sets (space M) (\<Union> (A ` {n..}))" by (auto simp: tail_events_def) |
|
69313 | 624 |
from this[of 0] have "X \<in> sigma_sets (space M) (\<Union>(A ` UNIV))" by simp |
42983 | 625 |
then show "X \<in> events" |
42982 | 626 |
by induct (insert A, auto) |
627 |
qed |
|
628 |
||
49772 | 629 |
lemma (in prob_space) sigma_algebra_tail_events: |
47694 | 630 |
assumes "\<And>i::nat. sigma_algebra (space M) (A i)" |
49772 | 631 |
shows "sigma_algebra (space M) (tail_events A)" |
632 |
unfolding tail_events_def |
|
42982 | 633 |
proof (simp add: sigma_algebra_iff2, safe) |
69325 | 634 |
let ?A = "(\<Inter>n. sigma_sets (space M) (\<Union> (A ` {n..})))" |
47694 | 635 |
interpret A: sigma_algebra "space M" "A i" for i by fact |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
636 |
{ fix X x assume "X \<in> ?A" "x \<in> X" |
69325 | 637 |
then have "\<And>n. X \<in> sigma_sets (space M) (\<Union> (A ` {n..}))" by auto |
69313 | 638 |
from this[of 0] have "X \<in> sigma_sets (space M) (\<Union>(A ` UNIV))" by simp |
42982 | 639 |
then have "X \<subseteq> space M" |
640 |
by induct (insert A.sets_into_space, auto) |
|
61808 | 641 |
with \<open>x \<in> X\<close> show "x \<in> space M" by auto } |
42982 | 642 |
{ fix F :: "nat \<Rightarrow> 'a set" and n assume "range F \<subseteq> ?A" |
69325 | 643 |
then show "(\<Union>(F ` UNIV)) \<in> sigma_sets (space M) (\<Union> (A ` {n..}))" |
42982 | 644 |
by (intro sigma_sets.Union) auto } |
645 |
qed (auto intro!: sigma_sets.Compl sigma_sets.Empty) |
|
646 |
||
647 |
lemma (in prob_space) kolmogorov_0_1_law: |
|
648 |
fixes A :: "nat \<Rightarrow> 'a set set" |
|
47694 | 649 |
assumes "\<And>i::nat. sigma_algebra (space M) (A i)" |
42982 | 650 |
assumes indep: "indep_sets A UNIV" |
49772 | 651 |
and X: "X \<in> tail_events A" |
42982 | 652 |
shows "prob X = 0 \<or> prob X = 1" |
653 |
proof - |
|
49781 | 654 |
have A: "\<And>i. A i \<subseteq> events" |
655 |
using indep unfolding indep_sets_def by simp |
|
656 |
||
47694 | 657 |
let ?D = "{D \<in> events. prob (X \<inter> D) = prob X * prob D}" |
658 |
interpret A: sigma_algebra "space M" "A i" for i by fact |
|
49772 | 659 |
interpret T: sigma_algebra "space M" "tail_events A" |
660 |
by (rule sigma_algebra_tail_events) fact |
|
42982 | 661 |
have "X \<subseteq> space M" using T.space_closed X by auto |
662 |
||
42983 | 663 |
have X_in: "X \<in> events" |
49772 | 664 |
using tail_events_sets A X by auto |
42982 | 665 |
|
69555 | 666 |
interpret D: Dynkin_system "space M" ?D |
667 |
proof (rule Dynkin_systemI) |
|
47694 | 668 |
fix D assume "D \<in> ?D" then show "D \<subseteq> space M" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
669 |
using sets.sets_into_space by auto |
42982 | 670 |
next |
47694 | 671 |
show "space M \<in> ?D" |
61808 | 672 |
using prob_space \<open>X \<subseteq> space M\<close> by (simp add: Int_absorb2) |
42982 | 673 |
next |
47694 | 674 |
fix A assume A: "A \<in> ?D" |
42982 | 675 |
have "prob (X \<inter> (space M - A)) = prob (X - (X \<inter> A))" |
61808 | 676 |
using \<open>X \<subseteq> space M\<close> by (auto intro!: arg_cong[where f=prob]) |
42982 | 677 |
also have "\<dots> = prob X - prob (X \<inter> A)" |
678 |
using X_in A by (intro finite_measure_Diff) auto |
|
679 |
also have "\<dots> = prob X * prob (space M) - prob X * prob A" |
|
680 |
using A prob_space by auto |
|
681 |
also have "\<dots> = prob X * prob (space M - A)" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
682 |
using X_in A sets.sets_into_space |
42982 | 683 |
by (subst finite_measure_Diff) (auto simp: field_simps) |
47694 | 684 |
finally show "space M - A \<in> ?D" |
61808 | 685 |
using A \<open>X \<subseteq> space M\<close> by auto |
42982 | 686 |
next |
47694 | 687 |
fix F :: "nat \<Rightarrow> 'a set" assume dis: "disjoint_family F" and "range F \<subseteq> ?D" |
42982 | 688 |
then have F: "range F \<subseteq> events" "\<And>i. prob (X \<inter> F i) = prob X * prob (F i)" |
689 |
by auto |
|
690 |
have "(\<lambda>i. prob (X \<inter> F i)) sums prob (\<Union>i. X \<inter> F i)" |
|
691 |
proof (rule finite_measure_UNION) |
|
692 |
show "range (\<lambda>i. X \<inter> F i) \<subseteq> events" |
|
693 |
using F X_in by auto |
|
694 |
show "disjoint_family (\<lambda>i. X \<inter> F i)" |
|
695 |
using dis by (rule disjoint_family_on_bisimulation) auto |
|
696 |
qed |
|
697 |
with F have "(\<lambda>i. prob X * prob (F i)) sums prob (X \<inter> (\<Union>i. F i))" |
|
698 |
by simp |
|
699 |
moreover have "(\<lambda>i. prob X * prob (F i)) sums (prob X * prob (\<Union>i. F i))" |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
43920
diff
changeset
|
700 |
by (intro sums_mult finite_measure_UNION F dis) |
42982 | 701 |
ultimately have "prob (X \<inter> (\<Union>i. F i)) = prob X * prob (\<Union>i. F i)" |
702 |
by (auto dest!: sums_unique) |
|
47694 | 703 |
with F show "(\<Union>i. F i) \<in> ?D" |
42982 | 704 |
by auto |
705 |
qed |
|
706 |
||
707 |
{ fix n |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
708 |
have "indep_sets (\<lambda>b. sigma_sets (space M) (\<Union>m\<in>case_bool {..n} {Suc n..} b. A m)) UNIV" |
42982 | 709 |
proof (rule indep_sets_collect_sigma) |
710 |
have *: "(\<Union>b. case b of True \<Rightarrow> {..n} | False \<Rightarrow> {Suc n..}) = UNIV" (is "?U = _") |
|
711 |
by (simp split: bool.split add: set_eq_iff) (metis not_less_eq_eq) |
|
712 |
with indep show "indep_sets A ?U" by simp |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
713 |
show "disjoint_family (case_bool {..n} {Suc n..})" |
42982 | 714 |
unfolding disjoint_family_on_def by (auto split: bool.split) |
715 |
fix m |
|
47694 | 716 |
show "Int_stable (A m)" |
42982 | 717 |
unfolding Int_stable_def using A.Int by auto |
718 |
qed |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
719 |
also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>case_bool {..n} {Suc n..} b. A m)) = |
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
720 |
case_bool (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))" |
42982 | 721 |
by (auto intro!: ext split: bool.split) |
722 |
finally have indep: "indep_set (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))" |
|
723 |
unfolding indep_set_def by simp |
|
724 |
||
47694 | 725 |
have "sigma_sets (space M) (\<Union>m\<in>{..n}. A m) \<subseteq> ?D" |
42982 | 726 |
proof (simp add: subset_eq, rule) |
727 |
fix D assume D: "D \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)" |
|
728 |
have "X \<in> sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m)" |
|
49772 | 729 |
using X unfolding tail_events_def by simp |
42982 | 730 |
from indep_setD[OF indep D this] indep_setD_ev1[OF indep] D |
731 |
show "D \<in> events \<and> prob (X \<inter> D) = prob X * prob D" |
|
732 |
by (auto simp add: ac_simps) |
|
733 |
qed } |
|
47694 | 734 |
then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> ?D" (is "?A \<subseteq> _") |
42982 | 735 |
by auto |
736 |
||
61808 | 737 |
note \<open>X \<in> tail_events A\<close> |
47694 | 738 |
also { |
739 |
have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> sigma_sets (space M) ?A" |
|
740 |
by (intro sigma_sets_subseteq UN_mono) auto |
|
49772 | 741 |
then have "tail_events A \<subseteq> sigma_sets (space M) ?A" |
742 |
unfolding tail_events_def by auto } |
|
69555 | 743 |
also have "sigma_sets (space M) ?A = Dynkin (space M) ?A" |
744 |
proof (rule sigma_eq_Dynkin) |
|
42982 | 745 |
{ fix B n assume "B \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)" |
746 |
then have "B \<subseteq> space M" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
747 |
by induct (insert A sets.sets_into_space[of _ M], auto) } |
47694 | 748 |
then show "?A \<subseteq> Pow (space M)" by auto |
749 |
show "Int_stable ?A" |
|
42982 | 750 |
proof (rule Int_stableI) |
74362 | 751 |
fix a b assume "a \<in> ?A" "b \<in> ?A" then obtain n m |
752 |
where a: "n \<in> UNIV" "a \<in> sigma_sets (space M) (\<Union> (A ` {..n}))" |
|
753 |
and b: "m \<in> UNIV" "b \<in> sigma_sets (space M) (\<Union> (A ` {..m}))" by auto |
|
47694 | 754 |
interpret Amn: sigma_algebra "space M" "sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
755 |
using A sets.sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto |
42982 | 756 |
have "sigma_sets (space M) (\<Union>i\<in>{..n}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" |
757 |
by (intro sigma_sets_subseteq UN_mono) auto |
|
758 |
with a have "a \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto |
|
759 |
moreover |
|
760 |
have "sigma_sets (space M) (\<Union>i\<in>{..m}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" |
|
761 |
by (intro sigma_sets_subseteq UN_mono) auto |
|
762 |
with b have "b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto |
|
763 |
ultimately have "a \<inter> b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" |
|
47694 | 764 |
using Amn.Int[of a b] by simp |
42982 | 765 |
then show "a \<inter> b \<in> (\<Union>n. sigma_sets (space M) (\<Union>i\<in>{..n}. A i))" by auto |
766 |
qed |
|
767 |
qed |
|
69555 | 768 |
also have "Dynkin (space M) ?A \<subseteq> ?D" |
769 |
using \<open>?A \<subseteq> ?D\<close> by (auto intro!: D.Dynkin_subset) |
|
47694 | 770 |
finally show ?thesis by auto |
42982 | 771 |
qed |
772 |
||
42985 | 773 |
lemma (in prob_space) borel_0_1_law: |
774 |
fixes F :: "nat \<Rightarrow> 'a set" |
|
49781 | 775 |
assumes F2: "indep_events F UNIV" |
42985 | 776 |
shows "prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 0 \<or> prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 1" |
777 |
proof (rule kolmogorov_0_1_law[of "\<lambda>i. sigma_sets (space M) { F i }"]) |
|
49781 | 778 |
have F1: "range F \<subseteq> events" |
779 |
using F2 by (simp add: indep_events_def subset_eq) |
|
47694 | 780 |
{ fix i show "sigma_algebra (space M) (sigma_sets (space M) {F i})" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
781 |
using sigma_algebra_sigma_sets[of "{F i}" "space M"] F1 sets.sets_into_space |
47694 | 782 |
by auto } |
42985 | 783 |
show "indep_sets (\<lambda>i. sigma_sets (space M) {F i}) UNIV" |
47694 | 784 |
proof (rule indep_sets_sigma) |
42985 | 785 |
show "indep_sets (\<lambda>i. {F i}) UNIV" |
49784
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents:
49781
diff
changeset
|
786 |
unfolding indep_events_def_alt[symmetric] by fact |
47694 | 787 |
fix i show "Int_stable {F i}" |
42985 | 788 |
unfolding Int_stable_def by simp |
789 |
qed |
|
46731 | 790 |
let ?Q = "\<lambda>n. \<Union>i\<in>{n..}. F i" |
49772 | 791 |
show "(\<Inter>n. \<Union>m\<in>{n..}. F m) \<in> tail_events (\<lambda>i. sigma_sets (space M) {F i})" |
792 |
unfolding tail_events_def |
|
42985 | 793 |
proof |
794 |
fix j |
|
47694 | 795 |
interpret S: sigma_algebra "space M" "sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
796 |
using order_trans[OF F1 sets.space_closed] |
47694 | 797 |
by (intro sigma_algebra_sigma_sets) (simp add: sigma_sets_singleton subset_eq) |
42985 | 798 |
have "(\<Inter>n. ?Q n) = (\<Inter>n\<in>{j..}. ?Q n)" |
799 |
by (intro decseq_SucI INT_decseq_offset UN_mono) auto |
|
47694 | 800 |
also have "\<dots> \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
801 |
using order_trans[OF F1 sets.space_closed] |
42985 | 802 |
by (safe intro!: S.countable_INT S.countable_UN) |
47694 | 803 |
(auto simp: sigma_sets_singleton intro!: sigma_sets.Basic bexI) |
42985 | 804 |
finally show "(\<Inter>n. ?Q n) \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})" |
47694 | 805 |
by simp |
42985 | 806 |
qed |
807 |
qed |
|
808 |
||
59000 | 809 |
lemma (in prob_space) borel_0_1_law_AE: |
810 |
fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool" |
|
811 |
assumes "indep_events (\<lambda>m. {x\<in>space M. P m x}) UNIV" (is "indep_events ?P _") |
|
812 |
shows "(AE x in M. infinite {m. P m x}) \<or> (AE x in M. finite {m. P m x})" |
|
813 |
proof - |
|
814 |
have [measurable]: "\<And>m. {x\<in>space M. P m x} \<in> sets M" |
|
815 |
using assms by (auto simp: indep_events_def) |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
816 |
have *: "(\<Inter>n. \<Union>m\<in>{n..}. {x \<in> space M. P m x}) \<in> events" |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
817 |
by simp |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
818 |
from assms have "prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 0 \<or> prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 1" |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
819 |
by (rule borel_0_1_law) |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
820 |
also have "prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 1 \<longleftrightarrow> (AE x in M. infinite {m. P m x})" |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
821 |
using * by (simp add: prob_eq_1) |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
822 |
(simp add: Bex_def infinite_nat_iff_unbounded_le) |
59000 | 823 |
also have "prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 0 \<longleftrightarrow> (AE x in M. finite {m. P m x})" |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
824 |
using * by (simp add: prob_eq_0) |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
825 |
(auto simp add: Ball_def finite_nat_iff_bounded not_less [symmetric]) |
59000 | 826 |
finally show ?thesis |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61808
diff
changeset
|
827 |
by blast |
59000 | 828 |
qed |
829 |
||
42987 | 830 |
lemma (in prob_space) indep_sets_finite: |
831 |
assumes I: "I \<noteq> {}" "finite I" |
|
832 |
and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> events" "\<And>i. i \<in> I \<Longrightarrow> space M \<in> F i" |
|
833 |
shows "indep_sets F I \<longleftrightarrow> (\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j)))" |
|
834 |
proof |
|
835 |
assume *: "indep_sets F I" |
|
836 |
from I show "\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j))" |
|
837 |
by (intro indep_setsD[OF *] ballI) auto |
|
838 |
next |
|
839 |
assume indep: "\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j))" |
|
840 |
show "indep_sets F I" |
|
841 |
proof (rule indep_setsI[OF F(1)]) |
|
842 |
fix A J assume J: "J \<noteq> {}" "J \<subseteq> I" "finite J" |
|
843 |
assume A: "\<forall>j\<in>J. A j \<in> F j" |
|
46731 | 844 |
let ?A = "\<lambda>j. if j \<in> J then A j else space M" |
42987 | 845 |
have "prob (\<Inter>j\<in>I. ?A j) = prob (\<Inter>j\<in>J. A j)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
846 |
using subset_trans[OF F(1) sets.space_closed] J A |
62390 | 847 |
by (auto intro!: arg_cong[where f=prob] split: if_split_asm) blast |
42987 | 848 |
also |
849 |
from A F have "(\<lambda>j. if j \<in> J then A j else space M) \<in> Pi I F" (is "?A \<in> _") |
|
62390 | 850 |
by (auto split: if_split_asm) |
42987 | 851 |
with indep have "prob (\<Inter>j\<in>I. ?A j) = (\<Prod>j\<in>I. prob (?A j))" |
852 |
by auto |
|
853 |
also have "\<dots> = (\<Prod>j\<in>J. prob (A j))" |
|
64272 | 854 |
unfolding if_distrib prod.If_cases[OF \<open>finite I\<close>] |
855 |
using prob_space \<open>J \<subseteq> I\<close> by (simp add: Int_absorb1 prod.neutral_const) |
|
42987 | 856 |
finally show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" .. |
857 |
qed |
|
858 |
qed |
|
859 |
||
42989 | 860 |
lemma (in prob_space) indep_vars_finite: |
42987 | 861 |
fixes I :: "'i set" |
862 |
assumes I: "I \<noteq> {}" "finite I" |
|
47694 | 863 |
and M': "\<And>i. i \<in> I \<Longrightarrow> sets (M' i) = sigma_sets (space (M' i)) (E i)" |
864 |
and rv: "\<And>i. i \<in> I \<Longrightarrow> random_variable (M' i) (X i)" |
|
865 |
and Int_stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (E i)" |
|
866 |
and space: "\<And>i. i \<in> I \<Longrightarrow> space (M' i) \<in> E i" and closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M' i))" |
|
867 |
shows "indep_vars M' X I \<longleftrightarrow> |
|
868 |
(\<forall>A\<in>(\<Pi> i\<in>I. E i). prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M)))" |
|
42987 | 869 |
proof - |
870 |
from rv have X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> space M \<rightarrow> space (M' i)" |
|
871 |
unfolding measurable_def by simp |
|
872 |
||
873 |
{ fix i assume "i\<in>I" |
|
61808 | 874 |
from closed[OF \<open>i \<in> I\<close>] |
47694 | 875 |
have "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)} |
876 |
= sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> E i}" |
|
61808 | 877 |
unfolding sigma_sets_vimage_commute[OF X, OF \<open>i \<in> I\<close>, symmetric] M'[OF \<open>i \<in> I\<close>] |
42987 | 878 |
by (subst sigma_sets_sigma_sets_eq) auto } |
47694 | 879 |
note sigma_sets_X = this |
42987 | 880 |
|
881 |
{ fix i assume "i\<in>I" |
|
47694 | 882 |
have "Int_stable {X i -` A \<inter> space M |A. A \<in> E i}" |
42987 | 883 |
proof (rule Int_stableI) |
47694 | 884 |
fix a assume "a \<in> {X i -` A \<inter> space M |A. A \<in> E i}" |
885 |
then obtain A where "a = X i -` A \<inter> space M" "A \<in> E i" by auto |
|
42987 | 886 |
moreover |
47694 | 887 |
fix b assume "b \<in> {X i -` A \<inter> space M |A. A \<in> E i}" |
888 |
then obtain B where "b = X i -` B \<inter> space M" "B \<in> E i" by auto |
|
42987 | 889 |
moreover |
890 |
have "(X i -` A \<inter> space M) \<inter> (X i -` B \<inter> space M) = X i -` (A \<inter> B) \<inter> space M" by auto |
|
61808 | 891 |
moreover note Int_stable[OF \<open>i \<in> I\<close>] |
42987 | 892 |
ultimately |
47694 | 893 |
show "a \<inter> b \<in> {X i -` A \<inter> space M |A. A \<in> E i}" |
42987 | 894 |
by (auto simp del: vimage_Int intro!: exI[of _ "A \<inter> B"] dest: Int_stableD) |
895 |
qed } |
|
47694 | 896 |
note indep_sets_X = indep_sets_sigma_sets_iff[OF this] |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
897 |
|
42987 | 898 |
{ fix i assume "i \<in> I" |
47694 | 899 |
{ fix A assume "A \<in> E i" |
61808 | 900 |
with M'[OF \<open>i \<in> I\<close>] have "A \<in> sets (M' i)" by auto |
42987 | 901 |
moreover |
61808 | 902 |
from rv[OF \<open>i\<in>I\<close>] have "X i \<in> measurable M (M' i)" by auto |
42987 | 903 |
ultimately |
904 |
have "X i -` A \<inter> space M \<in> sets M" by (auto intro: measurable_sets) } |
|
61808 | 905 |
with X[OF \<open>i\<in>I\<close>] space[OF \<open>i\<in>I\<close>] |
47694 | 906 |
have "{X i -` A \<inter> space M |A. A \<in> E i} \<subseteq> events" |
907 |
"space M \<in> {X i -` A \<inter> space M |A. A \<in> E i}" |
|
42987 | 908 |
by (auto intro!: exI[of _ "space (M' i)"]) } |
47694 | 909 |
note indep_sets_finite_X = indep_sets_finite[OF I this] |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
910 |
|
69313 | 911 |
have "(\<forall>A\<in>\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> E i}. prob (\<Inter>(A ` I)) = (\<Prod>j\<in>I. prob (A j))) = |
47694 | 912 |
(\<forall>A\<in>\<Pi> i\<in>I. E i. prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M)))" |
42987 | 913 |
(is "?L = ?R") |
914 |
proof safe |
|
47694 | 915 |
fix A assume ?L and A: "A \<in> (\<Pi> i\<in>I. E i)" |
61808 | 916 |
from \<open>?L\<close>[THEN bspec, of "\<lambda>i. X i -` A i \<inter> space M"] A \<open>I \<noteq> {}\<close> |
42987 | 917 |
show "prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M))" |
918 |
by (auto simp add: Pi_iff) |
|
919 |
next |
|
47694 | 920 |
fix A assume ?R and A: "A \<in> (\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> E i})" |
921 |
from A have "\<forall>i\<in>I. \<exists>B. A i = X i -` B \<inter> space M \<and> B \<in> E i" by auto |
|
42987 | 922 |
from bchoice[OF this] obtain B where B: "\<forall>i\<in>I. A i = X i -` B i \<inter> space M" |
47694 | 923 |
"B \<in> (\<Pi> i\<in>I. E i)" by auto |
61808 | 924 |
from \<open>?R\<close>[THEN bspec, OF B(2)] B(1) \<open>I \<noteq> {}\<close> |
69313 | 925 |
show "prob (\<Inter>(A ` I)) = (\<Prod>j\<in>I. prob (A j))" |
42987 | 926 |
by simp |
927 |
qed |
|
61808 | 928 |
then show ?thesis using \<open>I \<noteq> {}\<close> |
47694 | 929 |
by (simp add: rv indep_vars_def indep_sets_X sigma_sets_X indep_sets_finite_X cong: indep_sets_cong) |
42988 | 930 |
qed |
931 |
||
42989 | 932 |
lemma (in prob_space) indep_vars_compose: |
933 |
assumes "indep_vars M' X I" |
|
47694 | 934 |
assumes rv: "\<And>i. i \<in> I \<Longrightarrow> Y i \<in> measurable (M' i) (N i)" |
42989 | 935 |
shows "indep_vars N (\<lambda>i. Y i \<circ> X i) I" |
936 |
unfolding indep_vars_def |
|
42988 | 937 |
proof |
61808 | 938 |
from rv \<open>indep_vars M' X I\<close> |
42988 | 939 |
show "\<forall>i\<in>I. random_variable (N i) (Y i \<circ> X i)" |
47694 | 940 |
by (auto simp: indep_vars_def) |
42988 | 941 |
|
942 |
have "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I" |
|
61808 | 943 |
using \<open>indep_vars M' X I\<close> by (simp add: indep_vars_def) |
42988 | 944 |
then show "indep_sets (\<lambda>i. sigma_sets (space M) {(Y i \<circ> X i) -` A \<inter> space M |A. A \<in> sets (N i)}) I" |
945 |
proof (rule indep_sets_mono_sets) |
|
946 |
fix i assume "i \<in> I" |
|
61808 | 947 |
with \<open>indep_vars M' X I\<close> have X: "X i \<in> space M \<rightarrow> space (M' i)" |
42989 | 948 |
unfolding indep_vars_def measurable_def by auto |
42988 | 949 |
{ fix A assume "A \<in> sets (N i)" |
950 |
then have "\<exists>B. (Y i \<circ> X i) -` A \<inter> space M = X i -` B \<inter> space M \<and> B \<in> sets (M' i)" |
|
951 |
by (intro exI[of _ "Y i -` A \<inter> space (M' i)"]) |
|
61808 | 952 |
(auto simp: vimage_comp intro!: measurable_sets rv \<open>i \<in> I\<close> funcset_mem[OF X]) } |
42988 | 953 |
then show "sigma_sets (space M) {(Y i \<circ> X i) -` A \<inter> space M |A. A \<in> sets (N i)} \<subseteq> |
954 |
sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}" |
|
56154
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
55414
diff
changeset
|
955 |
by (intro sigma_sets_subseteq) (auto simp: vimage_comp) |
42988 | 956 |
qed |
957 |
qed |
|
958 |
||
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
959 |
lemma (in prob_space) indep_vars_compose2: |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
960 |
assumes "indep_vars M' X I" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
961 |
assumes rv: "\<And>i. i \<in> I \<Longrightarrow> Y i \<in> measurable (M' i) (N i)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
962 |
shows "indep_vars N (\<lambda>i x. Y i (X i x)) I" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
963 |
using indep_vars_compose [OF assms] by (simp add: comp_def) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
964 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
965 |
lemma (in prob_space) indep_var_compose: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
966 |
assumes "indep_var M1 X1 M2 X2" "Y1 \<in> measurable M1 N1" "Y2 \<in> measurable M2 N2" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
967 |
shows "indep_var N1 (Y1 \<circ> X1) N2 (Y2 \<circ> X2)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
968 |
proof - |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
969 |
have "indep_vars (case_bool N1 N2) (\<lambda>b. case_bool Y1 Y2 b \<circ> case_bool X1 X2 b) UNIV" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
970 |
using assms |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
971 |
by (intro indep_vars_compose[where M'="case_bool M1 M2"]) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
972 |
(auto simp: indep_var_def split: bool.split) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
973 |
also have "(\<lambda>b. case_bool Y1 Y2 b \<circ> case_bool X1 X2 b) = case_bool (Y1 \<circ> X1) (Y2 \<circ> X2)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
974 |
by (simp add: fun_eq_iff split: bool.split) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
975 |
finally show ?thesis |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
976 |
unfolding indep_var_def . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
977 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
978 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
979 |
lemma (in prob_space) indep_vars_Min: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
980 |
fixes X :: "'i \<Rightarrow> 'a \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
981 |
assumes I: "finite I" "i \<notin> I" and indep: "indep_vars (\<lambda>_. borel) X (insert i I)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
982 |
shows "indep_var borel (X i) borel (\<lambda>\<omega>. Min ((\<lambda>i. X i \<omega>)`I))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
983 |
proof - |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
984 |
have "indep_var |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
985 |
borel ((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i})) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
986 |
borel ((\<lambda>f. Min (f`I)) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
987 |
using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] borel_measurable_Min) auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
988 |
also have "((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i})) = X i" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
989 |
by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
990 |
also have "((\<lambda>f. Min (f`I)) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I)) = (\<lambda>\<omega>. Min ((\<lambda>i. X i \<omega>)`I))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
991 |
by (auto cong: rev_conj_cong) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
992 |
finally show ?thesis |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
993 |
unfolding indep_var_def . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
994 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
995 |
|
64267 | 996 |
lemma (in prob_space) indep_vars_sum: |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
997 |
fixes X :: "'i \<Rightarrow> 'a \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
998 |
assumes I: "finite I" "i \<notin> I" and indep: "indep_vars (\<lambda>_. borel) X (insert i I)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
999 |
shows "indep_var borel (X i) borel (\<lambda>\<omega>. \<Sum>i\<in>I. X i \<omega>)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1000 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1001 |
have "indep_var |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1002 |
borel ((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i})) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1003 |
borel ((\<lambda>f. \<Sum>i\<in>I. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1004 |
using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] ) auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1005 |
also have "((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i})) = X i" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1006 |
by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1007 |
also have "((\<lambda>f. \<Sum>i\<in>I. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I)) = (\<lambda>\<omega>. \<Sum>i\<in>I. X i \<omega>)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1008 |
by (auto cong: rev_conj_cong) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1009 |
finally show ?thesis . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1010 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1011 |
|
64272 | 1012 |
lemma (in prob_space) indep_vars_prod: |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1013 |
fixes X :: "'i \<Rightarrow> 'a \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1014 |
assumes I: "finite I" "i \<notin> I" and indep: "indep_vars (\<lambda>_. borel) X (insert i I)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1015 |
shows "indep_var borel (X i) borel (\<lambda>\<omega>. \<Prod>i\<in>I. X i \<omega>)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1016 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1017 |
have "indep_var |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1018 |
borel ((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i})) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1019 |
borel ((\<lambda>f. \<Prod>i\<in>I. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1020 |
using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] ) auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1021 |
also have "((\<lambda>f. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) {i})) = X i" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1022 |
by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1023 |
also have "((\<lambda>f. \<Prod>i\<in>I. f i) \<circ> (\<lambda>\<omega>. restrict (\<lambda>i. X i \<omega>) I)) = (\<lambda>\<omega>. \<Prod>i\<in>I. X i \<omega>)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1024 |
by (auto cong: rev_conj_cong) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1025 |
finally show ?thesis . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1026 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1027 |
|
47694 | 1028 |
lemma (in prob_space) indep_varsD_finite: |
42989 | 1029 |
assumes X: "indep_vars M' X I" |
42988 | 1030 |
assumes I: "I \<noteq> {}" "finite I" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M' i)" |
1031 |
shows "prob (\<Inter>i\<in>I. X i -` A i \<inter> space M) = (\<Prod>i\<in>I. prob (X i -` A i \<inter> space M))" |
|
1032 |
proof (rule indep_setsD) |
|
1033 |
show "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I" |
|
42989 | 1034 |
using X by (auto simp: indep_vars_def) |
42988 | 1035 |
show "I \<subseteq> I" "I \<noteq> {}" "finite I" using I by auto |
1036 |
show "\<forall>i\<in>I. X i -` A i \<inter> space M \<in> sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}" |
|
47694 | 1037 |
using I by auto |
42988 | 1038 |
qed |
1039 |
||
47694 | 1040 |
lemma (in prob_space) indep_varsD: |
1041 |
assumes X: "indep_vars M' X I" |
|
1042 |
assumes I: "J \<noteq> {}" "finite J" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M' i)" |
|
1043 |
shows "prob (\<Inter>i\<in>J. X i -` A i \<inter> space M) = (\<Prod>i\<in>J. prob (X i -` A i \<inter> space M))" |
|
1044 |
proof (rule indep_setsD) |
|
1045 |
show "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I" |
|
1046 |
using X by (auto simp: indep_vars_def) |
|
1047 |
show "\<forall>i\<in>J. X i -` A i \<inter> space M \<in> sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}" |
|
1048 |
using I by auto |
|
1049 |
qed fact+ |
|
1050 |
||
1051 |
lemma (in prob_space) indep_vars_iff_distr_eq_PiM: |
|
1052 |
fixes I :: "'i set" and X :: "'i \<Rightarrow> 'a \<Rightarrow> 'b" |
|
1053 |
assumes "I \<noteq> {}" |
|
42988 | 1054 |
assumes rv: "\<And>i. random_variable (M' i) (X i)" |
42989 | 1055 |
shows "indep_vars M' X I \<longleftrightarrow> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1056 |
distr M (\<Pi>\<^sub>M i\<in>I. M' i) (\<lambda>x. \<lambda>i\<in>I. X i x) = (\<Pi>\<^sub>M i\<in>I. distr M (M' i) (X i))" |
42988 | 1057 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1058 |
let ?P = "\<Pi>\<^sub>M i\<in>I. M' i" |
47694 | 1059 |
let ?X = "\<lambda>x. \<lambda>i\<in>I. X i x" |
1060 |
let ?D = "distr M ?P ?X" |
|
1061 |
have X: "random_variable ?P ?X" by (intro measurable_restrict rv) |
|
1062 |
interpret D: prob_space ?D by (intro prob_space_distr X) |
|
42988 | 1063 |
|
47694 | 1064 |
let ?D' = "\<lambda>i. distr M (M' i) (X i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1065 |
let ?P' = "\<Pi>\<^sub>M i\<in>I. distr M (M' i) (X i)" |
47694 | 1066 |
interpret D': prob_space "?D' i" for i by (intro prob_space_distr rv) |
1067 |
interpret P: product_prob_space ?D' I .. |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1068 |
|
42988 | 1069 |
show ?thesis |
47694 | 1070 |
proof |
42989 | 1071 |
assume "indep_vars M' X I" |
47694 | 1072 |
show "?D = ?P'" |
1073 |
proof (rule measure_eqI_generator_eq) |
|
1074 |
show "Int_stable (prod_algebra I M')" |
|
1075 |
by (rule Int_stable_prod_algebra) |
|
1076 |
show "prod_algebra I M' \<subseteq> Pow (space ?P)" |
|
1077 |
using prod_algebra_sets_into_space by (simp add: space_PiM) |
|
1078 |
show "sets ?D = sigma_sets (space ?P) (prod_algebra I M')" |
|
1079 |
by (simp add: sets_PiM space_PiM) |
|
1080 |
show "sets ?P' = sigma_sets (space ?P) (prod_algebra I M')" |
|
1081 |
by (simp add: sets_PiM space_PiM cong: prod_algebra_cong) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1082 |
let ?A = "\<lambda>i. \<Pi>\<^sub>E i\<in>I. space (M' i)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1083 |
show "range ?A \<subseteq> prod_algebra I M'" "(\<Union>i. ?A i) = space (Pi\<^sub>M I M')" |
47694 | 1084 |
by (auto simp: space_PiM intro!: space_in_prod_algebra cong: prod_algebra_cong) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1085 |
{ fix i show "emeasure ?D (\<Pi>\<^sub>E i\<in>I. space (M' i)) \<noteq> \<infinity>" by auto } |
47694 | 1086 |
next |
1087 |
fix E assume E: "E \<in> prod_algebra I M'" |
|
74362 | 1088 |
from prod_algebraE[OF E] obtain J Y |
1089 |
where J: |
|
1090 |
"E = prod_emb I M' J (Pi\<^sub>E J Y)" |
|
1091 |
"finite J" |
|
1092 |
"J \<noteq> {} \<or> I = {}" |
|
1093 |
"J \<subseteq> I" |
|
1094 |
"\<And>i. i \<in> J \<Longrightarrow> Y i \<in> sets (M' i)" |
|
1095 |
by auto |
|
47694 | 1096 |
from E have "E \<in> sets ?P" by (auto simp: sets_PiM) |
1097 |
then have "emeasure ?D E = emeasure M (?X -` E \<inter> space M)" |
|
1098 |
by (simp add: emeasure_distr X) |
|
1099 |
also have "?X -` E \<inter> space M = (\<Inter>i\<in>J. X i -` Y i \<inter> space M)" |
|
62390 | 1100 |
using J \<open>I \<noteq> {}\<close> measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: if_split_asm) |
47694 | 1101 |
also have "emeasure M (\<Inter>i\<in>J. X i -` Y i \<inter> space M) = (\<Prod> i\<in>J. emeasure M (X i -` Y i \<inter> space M))" |
61808 | 1102 |
using \<open>indep_vars M' X I\<close> J \<open>I \<noteq> {}\<close> using indep_varsD[of M' X I J] |
64272 | 1103 |
by (auto simp: emeasure_eq_measure prod_ennreal measure_nonneg prod_nonneg) |
47694 | 1104 |
also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))" |
1105 |
using rv J by (simp add: emeasure_distr) |
|
1106 |
also have "\<dots> = emeasure ?P' E" |
|
1107 |
using P.emeasure_PiM_emb[of J Y] J by (simp add: prod_emb_def) |
|
1108 |
finally show "emeasure ?D E = emeasure ?P' E" . |
|
42988 | 1109 |
qed |
1110 |
next |
|
47694 | 1111 |
assume "?D = ?P'" |
1112 |
show "indep_vars M' X I" unfolding indep_vars_def |
|
1113 |
proof (intro conjI indep_setsI ballI rv) |
|
1114 |
fix i show "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)} \<subseteq> events" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1115 |
by (auto intro!: sets.sigma_sets_subset measurable_sets rv) |
42988 | 1116 |
next |
47694 | 1117 |
fix J Y' assume J: "J \<noteq> {}" "J \<subseteq> I" "finite J" |
1118 |
assume Y': "\<forall>j\<in>J. Y' j \<in> sigma_sets (space M) {X j -` A \<inter> space M |A. A \<in> sets (M' j)}" |
|
1119 |
have "\<forall>j\<in>J. \<exists>Y. Y' j = X j -` Y \<inter> space M \<and> Y \<in> sets (M' j)" |
|
42988 | 1120 |
proof |
47694 | 1121 |
fix j assume "j \<in> J" |
1122 |
from Y'[rule_format, OF this] rv[of j] |
|
1123 |
show "\<exists>Y. Y' j = X j -` Y \<inter> space M \<and> Y \<in> sets (M' j)" |
|
1124 |
by (subst (asm) sigma_sets_vimage_commute[symmetric, of _ _ "space (M' j)"]) |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1125 |
(auto dest: measurable_space simp: sets.sigma_sets_eq) |
42988 | 1126 |
qed |
47694 | 1127 |
from bchoice[OF this] obtain Y where |
1128 |
Y: "\<And>j. j \<in> J \<Longrightarrow> Y' j = X j -` Y j \<inter> space M" "\<And>j. j \<in> J \<Longrightarrow> Y j \<in> sets (M' j)" by auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1129 |
let ?E = "prod_emb I M' J (Pi\<^sub>E J Y)" |
47694 | 1130 |
from Y have "(\<Inter>j\<in>J. Y' j) = ?X -` ?E \<inter> space M" |
62390 | 1131 |
using J \<open>I \<noteq> {}\<close> measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: if_split_asm) |
47694 | 1132 |
then have "emeasure M (\<Inter>j\<in>J. Y' j) = emeasure M (?X -` ?E \<inter> space M)" |
1133 |
by simp |
|
1134 |
also have "\<dots> = emeasure ?D ?E" |
|
1135 |
using Y J by (intro emeasure_distr[symmetric] X sets_PiM_I) auto |
|
1136 |
also have "\<dots> = emeasure ?P' ?E" |
|
61808 | 1137 |
using \<open>?D = ?P'\<close> by simp |
47694 | 1138 |
also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))" |
1139 |
using P.emeasure_PiM_emb[of J Y] J Y by (simp add: prod_emb_def) |
|
1140 |
also have "\<dots> = (\<Prod> i\<in>J. emeasure M (Y' i))" |
|
1141 |
using rv J Y by (simp add: emeasure_distr) |
|
1142 |
finally have "emeasure M (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. emeasure M (Y' i))" . |
|
1143 |
then show "prob (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. prob (Y' i))" |
|
64272 | 1144 |
by (auto simp: emeasure_eq_measure prod_ennreal measure_nonneg prod_nonneg) |
42988 | 1145 |
qed |
1146 |
qed |
|
42987 | 1147 |
qed |
1148 |
||
73253
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1149 |
lemma (in prob_space) indep_vars_iff_distr_eq_PiM': |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1150 |
fixes I :: "'i set" and X :: "'i \<Rightarrow> 'a \<Rightarrow> 'b" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1151 |
assumes "I \<noteq> {}" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1152 |
assumes rv: "\<And>i. i \<in> I \<Longrightarrow> random_variable (M' i) (X i)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1153 |
shows "indep_vars M' X I \<longleftrightarrow> |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1154 |
distr M (\<Pi>\<^sub>M i\<in>I. M' i) (\<lambda>x. \<lambda>i\<in>I. X i x) = (\<Pi>\<^sub>M i\<in>I. distr M (M' i) (X i))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1155 |
proof - |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1156 |
from assms obtain j where j: "j \<in> I" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1157 |
by auto |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1158 |
define N' where "N' = (\<lambda>i. if i \<in> I then M' i else M' j)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1159 |
define Y where "Y = (\<lambda>i. if i \<in> I then X i else X j)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1160 |
have rv: "random_variable (N' i) (Y i)" for i |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1161 |
using j by (auto simp: N'_def Y_def intro: assms) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1162 |
|
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1163 |
have "indep_vars M' X I = indep_vars N' Y I" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1164 |
by (intro indep_vars_cong) (auto simp: N'_def Y_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1165 |
also have "\<dots> \<longleftrightarrow> distr M (\<Pi>\<^sub>M i\<in>I. N' i) (\<lambda>x. \<lambda>i\<in>I. Y i x) = (\<Pi>\<^sub>M i\<in>I. distr M (N' i) (Y i))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1166 |
by (intro indep_vars_iff_distr_eq_PiM rv assms) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1167 |
also have "(\<Pi>\<^sub>M i\<in>I. N' i) = (\<Pi>\<^sub>M i\<in>I. M' i)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1168 |
by (intro PiM_cong) (simp_all add: N'_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1169 |
also have "(\<lambda>x. \<lambda>i\<in>I. Y i x) = (\<lambda>x. \<lambda>i\<in>I. X i x)" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1170 |
by (simp_all add: Y_def fun_eq_iff) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1171 |
also have "(\<Pi>\<^sub>M i\<in>I. distr M (N' i) (Y i)) = (\<Pi>\<^sub>M i\<in>I. distr M (M' i) (X i))" |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1172 |
by (intro PiM_cong distr_cong) (simp_all add: N'_def Y_def) |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1173 |
finally show ?thesis . |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1174 |
qed |
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents:
69555
diff
changeset
|
1175 |
|
42989 | 1176 |
lemma (in prob_space) indep_varD: |
1177 |
assumes indep: "indep_var Ma A Mb B" |
|
1178 |
assumes sets: "Xa \<in> sets Ma" "Xb \<in> sets Mb" |
|
1179 |
shows "prob ((\<lambda>x. (A x, B x)) -` (Xa \<times> Xb) \<inter> space M) = |
|
1180 |
prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)" |
|
1181 |
proof - |
|
1182 |
have "prob ((\<lambda>x. (A x, B x)) -` (Xa \<times> Xb) \<inter> space M) = |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
1183 |
prob (\<Inter>i\<in>UNIV. (case_bool A B i -` case_bool Xa Xb i \<inter> space M))" |
42989 | 1184 |
by (auto intro!: arg_cong[where f=prob] simp: UNIV_bool) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
1185 |
also have "\<dots> = (\<Prod>i\<in>UNIV. prob (case_bool A B i -` case_bool Xa Xb i \<inter> space M))" |
42989 | 1186 |
using indep unfolding indep_var_def |
1187 |
by (rule indep_varsD) (auto split: bool.split intro: sets) |
|
1188 |
also have "\<dots> = prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)" |
|
1189 |
unfolding UNIV_bool by simp |
|
1190 |
finally show ?thesis . |
|
1191 |
qed |
|
1192 |
||
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1193 |
lemma (in prob_space) prob_indep_random_variable: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1194 |
assumes ind[simp]: "indep_var N X N Y" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1195 |
assumes [simp]: "A \<in> sets N" "B \<in> sets N" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1196 |
shows "\<P>(x in M. X x \<in> A \<and> Y x \<in> B) = \<P>(x in M. X x \<in> A) * \<P>(x in M. Y x \<in> B)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1197 |
proof- |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1198 |
have " \<P>(x in M. (X x)\<in>A \<and> (Y x)\<in> B ) = prob ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1199 |
by (auto intro!: arg_cong[where f= prob]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1200 |
also have "...= prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1201 |
by (auto intro!: indep_varD[where Ma=N and Mb=N]) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1202 |
also have "... = \<P>(x in M. X x \<in> A) * \<P>(x in M. Y x \<in> B)" |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
67399
diff
changeset
|
1203 |
by (auto intro!: arg_cong2[where f= "(*)"] arg_cong[where f= prob]) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1204 |
finally show ?thesis . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1205 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1206 |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1207 |
lemma (in prob_space) |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1208 |
assumes "indep_var S X T Y" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1209 |
shows indep_var_rv1: "random_variable S X" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1210 |
and indep_var_rv2: "random_variable T Y" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1211 |
proof - |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53015
diff
changeset
|
1212 |
have "\<forall>i\<in>UNIV. random_variable (case_bool S T i) (case_bool X Y i)" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1213 |
using assms unfolding indep_var_def indep_vars_def by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1214 |
then show "random_variable S X" "random_variable T Y" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1215 |
unfolding UNIV_bool by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1216 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1217 |
|
47694 | 1218 |
lemma (in prob_space) indep_var_distribution_eq: |
1219 |
"indep_var S X T Y \<longleftrightarrow> random_variable S X \<and> random_variable T Y \<and> |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1220 |
distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" (is "_ \<longleftrightarrow> _ \<and> _ \<and> ?S \<Otimes>\<^sub>M ?T = ?J") |
47694 | 1221 |
proof safe |
1222 |
assume "indep_var S X T Y" |
|
1223 |
then show rvs: "random_variable S X" "random_variable T Y" |
|
1224 |
by (blast dest: indep_var_rv1 indep_var_rv2)+ |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1225 |
then have XY: "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
47694 | 1226 |
by (rule measurable_Pair) |
1227 |
||
1228 |
interpret X: prob_space ?S by (rule prob_space_distr) fact |
|
1229 |
interpret Y: prob_space ?T by (rule prob_space_distr) fact |
|
1230 |
interpret XY: pair_prob_space ?S ?T .. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1231 |
show "?S \<Otimes>\<^sub>M ?T = ?J" |
47694 | 1232 |
proof (rule pair_measure_eqI) |
1233 |
show "sigma_finite_measure ?S" .. |
|
1234 |
show "sigma_finite_measure ?T" .. |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1235 |
|
47694 | 1236 |
fix A B assume A: "A \<in> sets ?S" and B: "B \<in> sets ?T" |
1237 |
have "emeasure ?J (A \<times> B) = emeasure M ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)" |
|
1238 |
using A B by (intro emeasure_distr[OF XY]) auto |
|
1239 |
also have "\<dots> = emeasure M (X -` A \<inter> space M) * emeasure M (Y -` B \<inter> space M)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1240 |
using indep_varD[OF \<open>indep_var S X T Y\<close>, of A B] A B |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1241 |
by (simp add: emeasure_eq_measure measure_nonneg ennreal_mult) |
47694 | 1242 |
also have "\<dots> = emeasure ?S A * emeasure ?T B" |
1243 |
using rvs A B by (simp add: emeasure_distr) |
|
1244 |
finally show "emeasure ?S A * emeasure ?T B = emeasure ?J (A \<times> B)" by simp |
|
1245 |
qed simp |
|
1246 |
next |
|
1247 |
assume rvs: "random_variable S X" "random_variable T Y" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1248 |
then have XY: "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
47694 | 1249 |
by (rule measurable_Pair) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1250 |
|
47694 | 1251 |
let ?S = "distr M S X" and ?T = "distr M T Y" |
1252 |
interpret X: prob_space ?S by (rule prob_space_distr) fact |
|
1253 |
interpret Y: prob_space ?T by (rule prob_space_distr) fact |
|
1254 |
interpret XY: pair_prob_space ?S ?T .. |
|
1255 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1256 |
assume "?S \<Otimes>\<^sub>M ?T = ?J" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1257 |
|
47694 | 1258 |
{ fix S and X |
1259 |
have "Int_stable {X -` A \<inter> space M |A. A \<in> sets S}" |
|
1260 |
proof (safe intro!: Int_stableI) |
|
1261 |
fix A B assume "A \<in> sets S" "B \<in> sets S" |
|
1262 |
then show "\<exists>C. (X -` A \<inter> space M) \<inter> (X -` B \<inter> space M) = (X -` C \<inter> space M) \<and> C \<in> sets S" |
|
1263 |
by (intro exI[of _ "A \<inter> B"]) auto |
|
1264 |
qed } |
|
1265 |
note Int_stable = this |
|
1266 |
||
1267 |
show "indep_var S X T Y" unfolding indep_var_eq |
|
1268 |
proof (intro conjI indep_set_sigma_sets Int_stable rvs) |
|
1269 |
show "indep_set {X -` A \<inter> space M |A. A \<in> sets S} {Y -` A \<inter> space M |A. A \<in> sets T}" |
|
1270 |
proof (safe intro!: indep_setI) |
|
1271 |
{ fix A assume "A \<in> sets S" then show "X -` A \<inter> space M \<in> sets M" |
|
61808 | 1272 |
using \<open>X \<in> measurable M S\<close> by (auto intro: measurable_sets) } |
47694 | 1273 |
{ fix A assume "A \<in> sets T" then show "Y -` A \<inter> space M \<in> sets M" |
61808 | 1274 |
using \<open>Y \<in> measurable M T\<close> by (auto intro: measurable_sets) } |
47694 | 1275 |
next |
1276 |
fix A B assume ab: "A \<in> sets S" "B \<in> sets T" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1277 |
then have "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) = emeasure ?J (A \<times> B)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1278 |
using XY by (auto simp add: emeasure_distr emeasure_eq_measure measure_nonneg intro!: arg_cong[where f="prob"]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1279 |
also have "\<dots> = emeasure (?S \<Otimes>\<^sub>M ?T) (A \<times> B)" |
61808 | 1280 |
unfolding \<open>?S \<Otimes>\<^sub>M ?T = ?J\<close> .. |
47694 | 1281 |
also have "\<dots> = emeasure ?S A * emeasure ?T B" |
49776 | 1282 |
using ab by (simp add: Y.emeasure_pair_measure_Times) |
47694 | 1283 |
finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) = |
1284 |
prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1285 |
using rvs ab by (simp add: emeasure_eq_measure emeasure_distr measure_nonneg ennreal_mult[symmetric]) |
47694 | 1286 |
qed |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1287 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42989
diff
changeset
|
1288 |
qed |
42989 | 1289 |
|
49795 | 1290 |
lemma (in prob_space) distributed_joint_indep: |
1291 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
|
1292 |
assumes X: "distributed M S X Px" and Y: "distributed M T Y Py" |
|
1293 |
assumes indep: "indep_var S X T Y" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50244
diff
changeset
|
1294 |
shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)" |
49795 | 1295 |
using indep_var_distribution_eq[of S X T Y] indep |
1296 |
by (intro distributed_joint_indep'[OF S T X Y]) auto |
|
1297 |
||
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1298 |
lemma (in prob_space) indep_vars_nn_integral: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1299 |
assumes I: "finite I" "indep_vars (\<lambda>_. borel) X I" "\<And>i \<omega>. i \<in> I \<Longrightarrow> 0 \<le> X i \<omega>" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1300 |
shows "(\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<Prod>i\<in>I. \<integral>\<^sup>+\<omega>. X i \<omega> \<partial>M)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1301 |
proof cases |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1302 |
assume "I \<noteq> {}" |
63040 | 1303 |
define Y where [abs_def]: "Y i \<omega> = (if i \<in> I then X i \<omega> else 0)" for i \<omega> |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1304 |
{ fix i have "i \<in> I \<Longrightarrow> random_variable borel (X i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1305 |
using I(2) by (cases "i\<in>I") (auto simp: indep_vars_def) } |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1306 |
note rv_X = this |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1307 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1308 |
{ fix i have "random_variable borel (Y i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1309 |
using I(2) by (cases "i\<in>I") (auto simp: Y_def rv_X) } |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1310 |
note rv_Y = this[measurable] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1311 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1312 |
interpret Y: prob_space "distr M borel (Y i)" for i |
61359
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
59000
diff
changeset
|
1313 |
using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1314 |
interpret product_sigma_finite "\<lambda>i. distr M borel (Y i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1315 |
.. |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1316 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1317 |
have indep_Y: "indep_vars (\<lambda>i. borel) Y I" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1318 |
by (rule indep_vars_cong[THEN iffD1, OF _ _ _ I(2)]) (auto simp: Y_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1319 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1320 |
have "(\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. Y i \<omega>) \<partial>M)" |
64272 | 1321 |
using I(3) by (auto intro!: nn_integral_cong prod.cong simp add: Y_def max_def) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1322 |
also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x))" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1323 |
by (subst nn_integral_distr) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1324 |
also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>Pi\<^sub>M I (\<lambda>i. distr M borel (Y i)))" |
61808 | 1325 |
unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \<open>I \<noteq> {}\<close> rv_Y indep_Y] .. |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1326 |
also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^sup>+\<omega>. \<omega> \<partial>distr M borel (Y i)))" |
64272 | 1327 |
by (rule product_nn_integral_prod) (auto intro: \<open>finite I\<close>) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1328 |
also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+\<omega>. X i \<omega> \<partial>M)" |
64272 | 1329 |
by (intro prod.cong nn_integral_cong) (auto simp: nn_integral_distr Y_def rv_X) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1330 |
finally show ?thesis . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1331 |
qed (simp add: emeasure_space_1) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1332 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1333 |
lemma (in prob_space) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1334 |
fixes X :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{real_normed_field, banach, second_countable_topology}" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1335 |
assumes I: "finite I" "indep_vars (\<lambda>_. borel) X I" "\<And>i. i \<in> I \<Longrightarrow> integrable M (X i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1336 |
shows indep_vars_lebesgue_integral: "(\<integral>\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<Prod>i\<in>I. \<integral>\<omega>. X i \<omega> \<partial>M)" (is ?eq) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1337 |
and indep_vars_integrable: "integrable M (\<lambda>\<omega>. (\<Prod>i\<in>I. X i \<omega>))" (is ?int) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1338 |
proof (induct rule: case_split) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1339 |
assume "I \<noteq> {}" |
63040 | 1340 |
define Y where [abs_def]: "Y i \<omega> = (if i \<in> I then X i \<omega> else 0)" for i \<omega> |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1341 |
{ fix i have "i \<in> I \<Longrightarrow> random_variable borel (X i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1342 |
using I(2) by (cases "i\<in>I") (auto simp: indep_vars_def) } |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1343 |
note rv_X = this[measurable] |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1344 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1345 |
{ fix i have "random_variable borel (Y i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1346 |
using I(2) by (cases "i\<in>I") (auto simp: Y_def rv_X) } |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1347 |
note rv_Y = this[measurable] |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1348 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1349 |
{ fix i have "integrable M (Y i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1350 |
using I(3) by (cases "i\<in>I") (auto simp: Y_def) } |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1351 |
note int_Y = this |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1352 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1353 |
interpret Y: prob_space "distr M borel (Y i)" for i |
61359
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents:
59000
diff
changeset
|
1354 |
using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1355 |
interpret product_sigma_finite "\<lambda>i. distr M borel (Y i)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1356 |
.. |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1357 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1358 |
have indep_Y: "indep_vars (\<lambda>i. borel) Y I" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1359 |
by (rule indep_vars_cong[THEN iffD1, OF _ _ _ I(2)]) (auto simp: Y_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1360 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1361 |
have "(\<integral>\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<integral>\<omega>. (\<Prod>i\<in>I. Y i \<omega>) \<partial>M)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1362 |
using I(3) by (simp add: Y_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1363 |
also have "\<dots> = (\<integral>\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x))" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1364 |
by (subst integral_distr) auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1365 |
also have "\<dots> = (\<integral>\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>Pi\<^sub>M I (\<lambda>i. distr M borel (Y i)))" |
61808 | 1366 |
unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \<open>I \<noteq> {}\<close> rv_Y indep_Y] .. |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1367 |
also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<omega>. \<omega> \<partial>distr M borel (Y i)))" |
64272 | 1368 |
by (rule product_integral_prod) (auto intro: \<open>finite I\<close> simp: integrable_distr_eq int_Y) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1369 |
also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<omega>. X i \<omega> \<partial>M)" |
64272 | 1370 |
by (intro prod.cong integral_cong) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1371 |
(auto simp: integral_distr Y_def rv_X) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1372 |
finally show ?eq . |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1373 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1374 |
have "integrable (distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x)) (\<lambda>\<omega>. (\<Prod>i\<in>I. \<omega> i))" |
61808 | 1375 |
unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \<open>I \<noteq> {}\<close> rv_Y indep_Y] |
64272 | 1376 |
by (intro product_integrable_prod[OF \<open>finite I\<close>]) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1377 |
(simp add: integrable_distr_eq int_Y) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1378 |
then show ?int |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1379 |
by (simp add: integrable_distr_eq Y_def) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1380 |
qed (simp_all add: prob_space) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1381 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1382 |
lemma (in prob_space) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1383 |
fixes X1 X2 :: "'a \<Rightarrow> 'b::{real_normed_field, banach, second_countable_topology}" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1384 |
assumes "indep_var borel X1 borel X2" "integrable M X1" "integrable M X2" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1385 |
shows indep_var_lebesgue_integral: "(\<integral>\<omega>. X1 \<omega> * X2 \<omega> \<partial>M) = (\<integral>\<omega>. X1 \<omega> \<partial>M) * (\<integral>\<omega>. X2 \<omega> \<partial>M)" (is ?eq) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1386 |
and indep_var_integrable: "integrable M (\<lambda>\<omega>. X1 \<omega> * X2 \<omega>)" (is ?int) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1387 |
unfolding indep_var_def |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1388 |
proof - |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1389 |
have *: "(\<lambda>\<omega>. X1 \<omega> * X2 \<omega>) = (\<lambda>\<omega>. \<Prod>i\<in>UNIV. (case_bool X1 X2 i \<omega>))" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57447
diff
changeset
|
1390 |
by (simp add: UNIV_bool mult.commute) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1391 |
have **: "(\<lambda> _. borel) = case_bool borel borel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1392 |
by (rule ext, metis (full_types) bool.simps(3) bool.simps(4)) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1393 |
show ?eq |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1394 |
apply (subst *) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1395 |
apply (subst indep_vars_lebesgue_integral) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1396 |
apply (auto) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1397 |
apply (subst **, subst indep_var_def [symmetric], rule assms) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1398 |
apply (simp split: bool.split add: assms) |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57447
diff
changeset
|
1399 |
by (simp add: UNIV_bool mult.commute) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1400 |
show ?int |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1401 |
apply (subst *) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1402 |
apply (rule indep_vars_integrable) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1403 |
apply auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1404 |
apply (subst **, subst indep_var_def [symmetric], rule assms) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1405 |
by (simp split: bool.split add: assms) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1406 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
56154
diff
changeset
|
1407 |
|
42861
16375b493b64
Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff
changeset
|
1408 |
end |