| author | hoelzl | 
| Wed, 10 Oct 2012 12:12:28 +0200 | |
| changeset 49791 | e0854abfb3fc | 
| parent 49784 | 5e5b2da42a69 | 
| child 49794 | 3c7b5988e094 | 
| 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 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 3 | *) | 
| 
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 | header {* Independent families of events, event sets, and random variables *}
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 6 | |
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 7 | theory Independent_Family | 
| 47694 | 8 | imports Probability_Measure Infinite_Product_Measure | 
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 9 | begin | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 10 | |
| 42985 | 11 | lemma INT_decseq_offset: | 
| 12 | assumes "decseq F" | |
| 13 |   shows "(\<Inter>i. F i) = (\<Inter>i\<in>{n..}. F i)"
 | |
| 14 | proof safe | |
| 15 |   fix x i assume x: "x \<in> (\<Inter>i\<in>{n..}. F i)"
 | |
| 16 | show "x \<in> F i" | |
| 17 | proof cases | |
| 18 | from x have "x \<in> F n" by auto | |
| 19 | also assume "i \<le> n" with `decseq F` have "F n \<subseteq> F i" | |
| 20 | unfolding decseq_def by simp | |
| 21 | finally show ?thesis . | |
| 22 | qed (insert x, simp) | |
| 23 | qed auto | |
| 24 | ||
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 25 | definition (in prob_space) | 
| 42983 | 26 | "indep_sets F I \<longleftrightarrow> (\<forall>i\<in>I. F i \<subseteq> events) \<and> | 
| 42981 | 27 |     (\<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))))"
 | 
| 28 | ||
| 29 | definition (in prob_space) | |
| 30 | "indep_set A B \<longleftrightarrow> indep_sets (bool_case A B) UNIV" | |
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 31 | |
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 32 | definition (in prob_space) | 
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 33 |   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: 
49781diff
changeset | 34 | |
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 35 | lemma image_subset_iff_funcset: "F ` A \<subseteq> B \<longleftrightarrow> F \<in> A \<rightarrow> B" | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 36 | by auto | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 37 | |
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 38 | lemma (in prob_space) indep_events_def: | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 39 | "indep_events A I \<longleftrightarrow> (A`I \<subseteq> events) \<and> | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 40 |     (\<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: 
49781diff
changeset | 41 | unfolding indep_events_def_alt indep_sets_def | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 42 | apply (simp add: Ball_def Pi_iff image_subset_iff_funcset) | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 43 | apply (intro conj_cong refl arg_cong[where f=All] ext imp_cong) | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 44 | apply auto | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 45 | done | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 46 | |
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 47 | definition (in prob_space) | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 48 | "indep_event A B \<longleftrightarrow> indep_events (bool_case A B) UNIV" | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 49 | |
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 50 | definition (in prob_space) | 
| 42989 | 51 | "indep_vars M' X I \<longleftrightarrow> | 
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 52 | (\<forall>i\<in>I. random_variable (M' i) (X i)) \<and> | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 53 |     indep_sets (\<lambda>i. sigma_sets (space M) { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 54 | |
| 42989 | 55 | definition (in prob_space) | 
| 56 | "indep_var Ma A Mb B \<longleftrightarrow> indep_vars (bool_case Ma Mb) (bool_case A B) UNIV" | |
| 57 | ||
| 47694 | 58 | lemma (in prob_space) indep_sets_cong: | 
| 42981 | 59 | "I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i = G i) \<Longrightarrow> indep_sets F I \<longleftrightarrow> indep_sets G J" | 
| 60 | by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+ | |
| 61 | ||
| 62 | lemma (in prob_space) indep_events_finite_index_events: | |
| 63 |   "indep_events F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_events F J)"
 | |
| 64 | by (auto simp: indep_events_def) | |
| 65 | ||
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 66 | lemma (in prob_space) indep_sets_finite_index_sets: | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 67 |   "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 | 68 | proof (intro iffI allI impI) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 69 |   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 | 70 | show "indep_sets F I" unfolding indep_sets_def | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 71 | proof (intro conjI ballI allI impI) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 72 | fix i assume "i \<in> I" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 73 |     with *[THEN spec, of "{i}"] show "F i \<subseteq> events"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 74 | by (auto simp: indep_sets_def) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 75 | qed (insert *, auto simp: indep_sets_def) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 76 | qed (auto simp: indep_sets_def) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 77 | |
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 78 | lemma (in prob_space) indep_sets_mono_index: | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 79 | "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 | 80 | unfolding indep_sets_def by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 81 | |
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 82 | lemma (in prob_space) indep_sets_mono_sets: | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 83 | assumes indep: "indep_sets F I" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 84 | 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 | 85 | shows "indep_sets G I" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 86 | proof - | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 87 | 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 | 88 | using mono by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 89 | 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 | 90 | using mono by (auto simp: Pi_iff) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 91 | ultimately show ?thesis | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 92 | using indep by (auto simp: indep_sets_def) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 93 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 94 | |
| 49772 | 95 | lemma (in prob_space) indep_sets_mono: | 
| 96 | assumes indep: "indep_sets F I" | |
| 97 | assumes mono: "J \<subseteq> I" "\<And>i. i\<in>J \<Longrightarrow> G i \<subseteq> F i" | |
| 98 | shows "indep_sets G J" | |
| 99 | apply (rule indep_sets_mono_sets) | |
| 100 | apply (rule indep_sets_mono_index) | |
| 101 | apply (fact +) | |
| 102 | done | |
| 103 | ||
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 104 | lemma (in prob_space) indep_setsI: | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 105 | 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 | 106 |     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 | 107 | shows "indep_sets F I" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 108 | 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 | 109 | |
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 110 | lemma (in prob_space) indep_setsD: | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 111 |   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 | 112 | 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 | 113 | using assms unfolding indep_sets_def by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 114 | |
| 42982 | 115 | lemma (in prob_space) indep_setI: | 
| 116 | assumes ev: "A \<subseteq> events" "B \<subseteq> events" | |
| 117 | and indep: "\<And>a b. a \<in> A \<Longrightarrow> b \<in> B \<Longrightarrow> prob (a \<inter> b) = prob a * prob b" | |
| 118 | shows "indep_set A B" | |
| 119 | unfolding indep_set_def | |
| 120 | proof (rule indep_setsI) | |
| 121 |   fix F J assume "J \<noteq> {}" "J \<subseteq> UNIV"
 | |
| 122 | and F: "\<forall>j\<in>J. F j \<in> (case j of True \<Rightarrow> A | False \<Rightarrow> B)" | |
| 123 | have "J \<in> Pow UNIV" by auto | |
| 124 |   with F `J \<noteq> {}` indep[of "F True" "F False"]
 | |
| 125 | show "prob (\<Inter>j\<in>J. F j) = (\<Prod>j\<in>J. prob (F j))" | |
| 126 | unfolding UNIV_bool Pow_insert by (auto simp: ac_simps) | |
| 127 | qed (auto split: bool.split simp: ev) | |
| 128 | ||
| 129 | lemma (in prob_space) indep_setD: | |
| 130 | assumes indep: "indep_set A B" and ev: "a \<in> A" "b \<in> B" | |
| 131 | shows "prob (a \<inter> b) = prob a * prob b" | |
| 132 | using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "bool_case a b"] ev | |
| 133 | by (simp add: ac_simps UNIV_bool) | |
| 134 | ||
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 135 | lemma (in prob_space) indep_var_eq: | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 136 | "indep_var S X T Y \<longleftrightarrow> | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 137 | (random_variable S X \<and> random_variable T Y) \<and> | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 138 | indep_set | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 139 |       (sigma_sets (space M) { X -` A \<inter> space M | A. A \<in> sets S})
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 140 |       (sigma_sets (space M) { Y -` A \<inter> space M | A. A \<in> sets T})"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 141 | unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 142 | by (intro arg_cong2[where f="op \<and>"] arg_cong2[where f=indep_sets] ext) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 143 | (auto split: bool.split) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 144 | |
| 42982 | 145 | lemma (in prob_space) | 
| 146 | assumes indep: "indep_set A B" | |
| 42983 | 147 | shows indep_setD_ev1: "A \<subseteq> events" | 
| 148 | and indep_setD_ev2: "B \<subseteq> events" | |
| 42982 | 149 | using indep unfolding indep_set_def indep_sets_def UNIV_bool by auto | 
| 150 | ||
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 151 | lemma (in prob_space) indep_sets_dynkin: | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 152 | assumes indep: "indep_sets F I" | 
| 47694 | 153 | 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 | 154 | (is "indep_sets ?F I") | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 155 | 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 | 156 |   fix J assume "finite J" "J \<subseteq> I" "J \<noteq> {}"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 157 | with indep have "indep_sets F J" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 158 | by (subst (asm) indep_sets_finite_index_sets) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 159 |   { fix J K assume "indep_sets F K"
 | 
| 46731 | 160 | 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 | 161 | assume "finite J" "J \<subseteq> K" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 162 | then have "indep_sets (?G J) K" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 163 | proof induct | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 164 | case (insert j J) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 165 | moreover def G \<equiv> "?G J" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 166 | 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 | 167 | by (auto simp: indep_sets_def) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 168 |       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 | 169 |       { fix X assume X: "X \<in> events"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 170 |         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 | 171 | \<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 | 172 |         have "indep_sets (G(j := {X})) K"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 173 | proof (rule indep_setsI) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 174 |           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 | 175 | using G X by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 176 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 177 |           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 | 178 | 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 | 179 | proof cases | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 180 | assume "j \<in> J" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 181 | with J have "A j = X" by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 182 | show ?thesis | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 183 | proof cases | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 184 |               assume "J = {j}" then show ?thesis by simp
 | 
| 
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 |               assume "J \<noteq> {j}"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 187 |               have "prob (\<Inter>i\<in>J. A i) = prob ((\<Inter>i\<in>J-{j}. A i) \<inter> X)"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 188 | using `j \<in> J` `A j = X` by (auto intro!: arg_cong[where f=prob] split: split_if_asm) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 189 |               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 | 190 | proof (rule indep) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 191 |                 show "J - {j} \<noteq> {}" "J - {j} \<subseteq> K" "finite (J - {j})" "j \<notin> J - {j}"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 192 |                   using J `J \<noteq> {j}` `j \<in> J` by auto
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 193 |                 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 | 194 | using J by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 195 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 196 |               also have "\<dots> = prob (A j) * (\<Prod>i\<in>J-{j}. prob (A i))"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 197 | using `A j = X` by simp | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 198 | also have "\<dots> = (\<Prod>i\<in>J. prob (A i))" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 199 | unfolding setprod.insert_remove[OF `finite J`, symmetric, of "\<lambda>i. prob (A i)"] | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 200 | using `j \<in> J` by (simp add: insert_absorb) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 201 | finally show ?thesis . | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 202 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 203 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 204 | assume "j \<notin> J" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 205 | with J have "\<forall>i\<in>J. A i \<in> G i" by (auto split: split_if_asm) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 206 | with J show ?thesis | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 207 | by (intro indep_setsD[OF G(1)]) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 208 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 209 | qed } | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 210 | note indep_sets_insert = this | 
| 47694 | 211 | have "dynkin_system (space M) ?D" | 
| 42987 | 212 | 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 | 213 |         show "indep_sets (G(j := {{}})) K"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 214 | by (rule indep_sets_insert) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 215 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 216 |         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 | 217 |         show "indep_sets (G(j := {space M - X})) K"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 218 | proof (rule indep_sets_insert) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 219 |           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 | 220 | 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 | 221 | using G by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 222 | 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 | 223 | prob ((\<Inter>j\<in>J. A j) - (\<Inter>i\<in>insert j J. (A(j := X)) i))" | 
| 47694 | 224 |             using A_sets sets_into_space[of _ M] X `J \<noteq> {}`
 | 
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 225 | by (auto intro!: arg_cong[where f=prob] split: split_if_asm) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 226 | also have "\<dots> = 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 | 227 |             using J `J \<noteq> {}` `j \<notin> J` A_sets X sets_into_space
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 228 | by (auto intro!: finite_measure_Diff finite_INT split: split_if_asm) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 229 | 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 | 230 | 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 | 231 |           moreover {
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 232 | have "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 | 233 | using J A `finite J` by (intro indep_setsD[OF G(1)]) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 234 | 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 | 235 | using prob_space by simp } | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 236 |           moreover {
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 237 | have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = (\<Prod>i\<in>insert j J. prob ((A(j := X)) i))" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 238 | using J A `j \<in> K` by (intro indep_setsD[OF G']) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 239 | then have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = prob X * (\<Prod>i\<in>J. prob (A i))" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 240 | using `finite J` `j \<notin> J` by (auto intro!: setprod_cong) } | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 241 | 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 | 242 | by (simp add: field_simps) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 243 | 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 | 244 | using X A by (simp add: finite_measure_compl) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 245 | 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 | 246 | qed (insert X, auto) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 247 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 248 | 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 | 249 |         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 | 250 |         show "indep_sets (G(j := {\<Union>k. F k})) K"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 251 | proof (rule indep_sets_insert) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 252 |           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 | 253 | 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 | 254 | using G by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 255 | 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))" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 256 |             using `J \<noteq> {}` `j \<notin> J` `j \<in> K` by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 257 | 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 | 258 | proof (rule finite_measure_UNION) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 259 | 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 | 260 | using disj by (rule disjoint_family_on_bisimulation) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 261 | show "range (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i) \<subseteq> events" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 262 |               using A_sets F `finite J` `J \<noteq> {}` `j \<notin> J` by (auto intro!: Int)
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 263 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 264 |           moreover { fix k
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 265 | from J A `j \<in> K` have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * (\<Prod>i\<in>J. prob (A i))" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 266 | by (subst indep_setsD[OF F(2)]) (auto intro!: setprod_cong split: split_if_asm) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 267 | also have "\<dots> = prob (F k) * prob (\<Inter>i\<in>J. A i)" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 268 | using J A `j \<in> K` by (subst indep_setsD[OF G(1)]) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 269 | 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 | 270 | 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 | 271 | by simp | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 272 | moreover | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 273 | 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 | 274 | 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 | 275 | 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)))" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 276 | using J A `j \<in> K` by (subst indep_setsD[OF G(1), symmetric]) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 277 | ultimately | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 278 | 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 | 279 | by (auto dest!: sums_unique) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 280 | qed (insert F, auto) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 281 | qed (insert sets_into_space, auto) | 
| 47694 | 282 |       then have mono: "dynkin (space M) (G j) \<subseteq> {E \<in> events. indep_sets (G(j := {E})) K}"
 | 
| 283 | proof (rule dynkin_system.dynkin_subset, safe) | |
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 284 | fix X assume "X \<in> G j" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 285 | then show "X \<in> events" using G `j \<in> K` by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 286 | from `indep_sets G K` | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 287 |         show "indep_sets (G(j := {X})) K"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 288 | by (rule indep_sets_mono_sets) (insert `X \<in> G j`, auto) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 289 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 290 | have "indep_sets (G(j:=?D)) K" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 291 | proof (rule indep_setsI) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 292 | 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 | 293 | using G(2) by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 294 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 295 |         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 | 296 | 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 | 297 | proof cases | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 298 | assume "j \<in> J" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 299 |           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 | 300 | from J A show ?thesis | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 301 | by (intro indep_setsD[OF indep]) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 302 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 303 | assume "j \<notin> J" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 304 | with J A have "\<forall>i\<in>J. A i \<in> G i" by (auto split: split_if_asm) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 305 | with J show ?thesis | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 306 | by (intro indep_setsD[OF G(1)]) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 307 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 308 | qed | 
| 47694 | 309 | 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 | 310 | by (rule indep_sets_mono_sets) (insert mono, auto) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 311 | then show ?case | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 312 | by (rule indep_sets_mono_sets) (insert `j \<in> K` `j \<notin> J`, auto simp: G_def) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 313 | qed (insert `indep_sets F K`, simp) } | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 314 | from this[OF `indep_sets F J` `finite J` subset_refl] | 
| 47694 | 315 | show "indep_sets ?F J" | 
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 316 | by (rule indep_sets_mono_sets) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 317 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 318 | |
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 319 | lemma (in prob_space) indep_sets_sigma: | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 320 | assumes indep: "indep_sets F I" | 
| 47694 | 321 | assumes stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (F i)" | 
| 322 | 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 | 323 | proof - | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 324 | from indep_sets_dynkin[OF indep] | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 325 | show ?thesis | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 326 | proof (rule indep_sets_mono_sets, subst sigma_eq_dynkin, simp_all add: stable) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 327 | fix i assume "i \<in> I" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 328 | with indep have "F i \<subseteq> events" by (auto simp: indep_sets_def) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 329 | with sets_into_space show "F i \<subseteq> Pow (space M)" by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 330 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 331 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 332 | |
| 42987 | 333 | lemma (in prob_space) indep_sets_sigma_sets_iff: | 
| 47694 | 334 | assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable (F i)" | 
| 42987 | 335 | shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I \<longleftrightarrow> indep_sets F I" | 
| 336 | proof | |
| 337 | assume "indep_sets F I" then show "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I" | |
| 47694 | 338 | by (rule indep_sets_sigma) fact | 
| 42987 | 339 | next | 
| 340 | assume "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I" then show "indep_sets F I" | |
| 341 | by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic) | |
| 342 | qed | |
| 343 | ||
| 49781 | 344 | lemma (in prob_space) | 
| 345 | "indep_vars M' X I \<longleftrightarrow> | |
| 346 | (\<forall>i\<in>I. random_variable (M' i) (X i)) \<and> | |
| 347 |     indep_sets (\<lambda>i. { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
 | |
| 348 | unfolding indep_vars_def | |
| 349 | apply (rule conj_cong[OF refl]) | |
| 350 | apply (rule indep_sets_sigma_sets_iff) | |
| 351 | apply (auto simp: Int_stable_def) | |
| 352 | apply (rule_tac x="A \<inter> Aa" in exI) | |
| 353 | apply auto | |
| 354 | done | |
| 355 | ||
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 356 | lemma (in prob_space) indep_sets2_eq: | 
| 42981 | 357 | "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)" | 
| 358 | unfolding indep_set_def | |
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 359 | proof (intro iffI ballI conjI) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 360 | assume indep: "indep_sets (bool_case A B) UNIV" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 361 |   { fix a b assume "a \<in> A" "b \<in> B"
 | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 362 | with indep_setsD[OF indep, of UNIV "bool_case a b"] | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 363 | show "prob (a \<inter> b) = prob a * prob b" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 364 | unfolding UNIV_bool by (simp add: ac_simps) } | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 365 | from indep show "A \<subseteq> events" "B \<subseteq> events" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 366 | unfolding indep_sets_def UNIV_bool by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 367 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 368 | 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)" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 369 | show "indep_sets (bool_case A B) UNIV" | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 370 | proof (rule indep_setsI) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 371 | 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 | 372 | using * by (auto split: bool.split) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 373 | next | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 374 |     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 | 375 |     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 | 376 | by (auto simp: UNIV_bool) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 377 | 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 | 378 | using X * by auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 379 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 380 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 381 | |
| 42981 | 382 | lemma (in prob_space) indep_set_sigma_sets: | 
| 383 | assumes "indep_set A B" | |
| 47694 | 384 | assumes A: "Int_stable A" and B: "Int_stable B" | 
| 42981 | 385 | 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 | 386 | proof - | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 387 | have "indep_sets (\<lambda>i. sigma_sets (space M) (case i of True \<Rightarrow> A | False \<Rightarrow> B)) UNIV" | 
| 47694 | 388 | proof (rule indep_sets_sigma) | 
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 389 | show "indep_sets (bool_case A B) UNIV" | 
| 42981 | 390 | by (rule `indep_set A B`[unfolded indep_set_def]) | 
| 47694 | 391 | 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 | 392 | using A B by (cases i) auto | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 393 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 394 | then show ?thesis | 
| 42981 | 395 | unfolding indep_set_def | 
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 396 | by (rule indep_sets_mono_sets) (auto split: bool.split) | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 397 | qed | 
| 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 398 | |
| 42981 | 399 | lemma (in prob_space) indep_sets_collect_sigma: | 
| 400 | fixes I :: "'j \<Rightarrow> 'i set" and J :: "'j set" and E :: "'i \<Rightarrow> 'a set set" | |
| 401 | assumes indep: "indep_sets E (\<Union>j\<in>J. I j)" | |
| 47694 | 402 | assumes Int_stable: "\<And>i j. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> Int_stable (E i)" | 
| 42981 | 403 | assumes disjoint: "disjoint_family_on I J" | 
| 404 | shows "indep_sets (\<lambda>j. sigma_sets (space M) (\<Union>i\<in>I j. E i)) J" | |
| 405 | proof - | |
| 46731 | 406 |   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 | 407 | |
| 42983 | 408 | from indep have E: "\<And>j i. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> E i \<subseteq> events" | 
| 42981 | 409 | unfolding indep_sets_def by auto | 
| 410 |   { fix j
 | |
| 47694 | 411 | let ?S = "sigma_sets (space M) (\<Union>i\<in>I j. E i)" | 
| 42981 | 412 | assume "j \<in> J" | 
| 47694 | 413 | from E[OF this] interpret S: sigma_algebra "space M" ?S | 
| 414 | using sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto | |
| 42981 | 415 | |
| 416 | have "sigma_sets (space M) (\<Union>i\<in>I j. E i) = sigma_sets (space M) (?E j)" | |
| 417 | proof (rule sigma_sets_eqI) | |
| 418 | fix A assume "A \<in> (\<Union>i\<in>I j. E i)" | |
| 419 | then guess i .. | |
| 420 | then show "A \<in> sigma_sets (space M) (?E j)" | |
| 47694 | 421 |         by (auto intro!: sigma_sets.intros(2-) exI[of _ "{i}"] exI[of _ "\<lambda>i. A"])
 | 
| 42981 | 422 | next | 
| 423 | fix A assume "A \<in> ?E j" | |
| 424 |       then obtain E' K where "finite K" "K \<noteq> {}" "K \<subseteq> I j" "\<And>k. k \<in> K \<Longrightarrow> E' k \<in> E k"
 | |
| 425 | and A: "A = (\<Inter>k\<in>K. E' k)" | |
| 426 | by auto | |
| 47694 | 427 | then have "A \<in> ?S" unfolding A | 
| 428 | by (safe intro!: S.finite_INT) auto | |
| 42981 | 429 | then show "A \<in> sigma_sets (space M) (\<Union>i\<in>I j. E i)" | 
| 47694 | 430 | by simp | 
| 42981 | 431 | qed } | 
| 432 | moreover have "indep_sets (\<lambda>j. sigma_sets (space M) (?E j)) J" | |
| 47694 | 433 | proof (rule indep_sets_sigma) | 
| 42981 | 434 | show "indep_sets ?E J" | 
| 435 | proof (intro indep_setsI) | |
| 436 | fix j assume "j \<in> J" with E show "?E j \<subseteq> events" by (force intro!: finite_INT) | |
| 437 | next | |
| 438 |       fix K A assume K: "K \<noteq> {}" "K \<subseteq> J" "finite K"
 | |
| 439 | and "\<forall>j\<in>K. A j \<in> ?E j" | |
| 440 |       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)"
 | |
| 441 | by simp | |
| 442 | from bchoice[OF this] guess E' .. | |
| 443 | from bchoice[OF this] obtain L | |
| 444 | where A: "\<And>j. j\<in>K \<Longrightarrow> A j = (\<Inter>l\<in>L j. E' j l)" | |
| 445 |         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"
 | |
| 446 | and E': "\<And>j l. j\<in>K \<Longrightarrow> l \<in> L j \<Longrightarrow> E' j l \<in> E l" | |
| 447 | by auto | |
| 448 | ||
| 449 |       { fix k l j assume "k \<in> K" "j \<in> K" "l \<in> L j" "l \<in> L k"
 | |
| 450 | have "k = j" | |
| 451 | proof (rule ccontr) | |
| 452 | assume "k \<noteq> j" | |
| 453 |           with disjoint `K \<subseteq> J` `k \<in> K` `j \<in> K` have "I k \<inter> I j = {}"
 | |
| 454 | unfolding disjoint_family_on_def by auto | |
| 455 | with L(2,3)[OF `j \<in> K`] L(2,3)[OF `k \<in> K`] | |
| 456 | show False using `l \<in> L k` `l \<in> L j` by auto | |
| 457 | qed } | |
| 458 | note L_inj = this | |
| 459 | ||
| 460 | def k \<equiv> "\<lambda>l. (SOME k. k \<in> K \<and> l \<in> L k)" | |
| 461 |       { fix x j l assume *: "j \<in> K" "l \<in> L j"
 | |
| 462 | have "k l = j" unfolding k_def | |
| 463 | proof (rule some_equality) | |
| 464 | fix k assume "k \<in> K \<and> l \<in> L k" | |
| 465 | with * L_inj show "k = j" by auto | |
| 466 | qed (insert *, simp) } | |
| 467 | note k_simp[simp] = this | |
| 46731 | 468 | let ?E' = "\<lambda>l. E' (k l) l" | 
| 42981 | 469 | have "prob (\<Inter>j\<in>K. A j) = prob (\<Inter>l\<in>(\<Union>k\<in>K. L k). ?E' l)" | 
| 470 | by (auto simp: A intro!: arg_cong[where f=prob]) | |
| 471 | also have "\<dots> = (\<Prod>l\<in>(\<Union>k\<in>K. L k). prob (?E' l))" | |
| 472 | using L K E' by (intro indep_setsD[OF indep]) (simp_all add: UN_mono) | |
| 473 | also have "\<dots> = (\<Prod>j\<in>K. \<Prod>l\<in>L j. prob (E' j l))" | |
| 474 | using K L L_inj by (subst setprod_UN_disjoint) auto | |
| 475 | also have "\<dots> = (\<Prod>j\<in>K. prob (A j))" | |
| 476 | using K L E' by (auto simp add: A intro!: setprod_cong indep_setsD[OF indep, symmetric]) blast | |
| 477 | finally show "prob (\<Inter>j\<in>K. A j) = (\<Prod>j\<in>K. prob (A j))" . | |
| 478 | qed | |
| 479 | next | |
| 480 | fix j assume "j \<in> J" | |
| 47694 | 481 | show "Int_stable (?E j)" | 
| 42981 | 482 | proof (rule Int_stableI) | 
| 483 | fix a assume "a \<in> ?E j" then obtain Ka Ea | |
| 484 |         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
 | |
| 485 | fix b assume "b \<in> ?E j" then obtain Kb Eb | |
| 486 |         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
 | |
| 487 |       let ?A = "\<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 {})"
 | |
| 488 | have "a \<inter> b = INTER (Ka \<union> Kb) ?A" | |
| 489 | by (simp add: a b set_eq_iff) auto | |
| 490 | with a b `j \<in> J` Int_stableD[OF Int_stable] show "a \<inter> b \<in> ?E j" | |
| 491 | by (intro CollectI exI[of _ "Ka \<union> Kb"] exI[of _ ?A]) auto | |
| 492 | qed | |
| 493 | qed | |
| 494 | ultimately show ?thesis | |
| 495 | by (simp cong: indep_sets_cong) | |
| 496 | qed | |
| 497 | ||
| 49772 | 498 | definition (in prob_space) tail_events where | 
| 499 |   "tail_events A = (\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
 | |
| 42982 | 500 | |
| 49772 | 501 | lemma (in prob_space) tail_events_sets: | 
| 502 | assumes A: "\<And>i::nat. A i \<subseteq> events" | |
| 503 | shows "tail_events A \<subseteq> events" | |
| 504 | proof | |
| 505 | fix X assume X: "X \<in> tail_events A" | |
| 42982 | 506 |   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
 | 
| 49772 | 507 |   from X have "\<And>n::nat. X \<in> sigma_sets (space M) (UNION {n..} A)" by (auto simp: tail_events_def)
 | 
| 42982 | 508 | from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp | 
| 42983 | 509 | then show "X \<in> events" | 
| 42982 | 510 | by induct (insert A, auto) | 
| 511 | qed | |
| 512 | ||
| 49772 | 513 | lemma (in prob_space) sigma_algebra_tail_events: | 
| 47694 | 514 | assumes "\<And>i::nat. sigma_algebra (space M) (A i)" | 
| 49772 | 515 | shows "sigma_algebra (space M) (tail_events A)" | 
| 516 | unfolding tail_events_def | |
| 42982 | 517 | proof (simp add: sigma_algebra_iff2, safe) | 
| 518 |   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
 | |
| 47694 | 519 | interpret A: sigma_algebra "space M" "A i" for i by fact | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 520 |   { fix X x assume "X \<in> ?A" "x \<in> X"
 | 
| 42982 | 521 |     then have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by auto
 | 
| 522 | from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp | |
| 523 | then have "X \<subseteq> space M" | |
| 524 | by induct (insert A.sets_into_space, auto) | |
| 525 | with `x \<in> X` show "x \<in> space M" by auto } | |
| 526 |   { fix F :: "nat \<Rightarrow> 'a set" and n assume "range F \<subseteq> ?A"
 | |
| 527 |     then show "(UNION UNIV F) \<in> sigma_sets (space M) (UNION {n..} A)"
 | |
| 528 | by (intro sigma_sets.Union) auto } | |
| 529 | qed (auto intro!: sigma_sets.Compl sigma_sets.Empty) | |
| 530 | ||
| 531 | lemma (in prob_space) kolmogorov_0_1_law: | |
| 532 | fixes A :: "nat \<Rightarrow> 'a set set" | |
| 47694 | 533 | assumes "\<And>i::nat. sigma_algebra (space M) (A i)" | 
| 42982 | 534 | assumes indep: "indep_sets A UNIV" | 
| 49772 | 535 | and X: "X \<in> tail_events A" | 
| 42982 | 536 | shows "prob X = 0 \<or> prob X = 1" | 
| 537 | proof - | |
| 49781 | 538 | have A: "\<And>i. A i \<subseteq> events" | 
| 539 | using indep unfolding indep_sets_def by simp | |
| 540 | ||
| 47694 | 541 |   let ?D = "{D \<in> events. prob (X \<inter> D) = prob X * prob D}"
 | 
| 542 | interpret A: sigma_algebra "space M" "A i" for i by fact | |
| 49772 | 543 | interpret T: sigma_algebra "space M" "tail_events A" | 
| 544 | by (rule sigma_algebra_tail_events) fact | |
| 42982 | 545 | have "X \<subseteq> space M" using T.space_closed X by auto | 
| 546 | ||
| 42983 | 547 | have X_in: "X \<in> events" | 
| 49772 | 548 | using tail_events_sets A X by auto | 
| 42982 | 549 | |
| 47694 | 550 | interpret D: dynkin_system "space M" ?D | 
| 42982 | 551 | proof (rule dynkin_systemI) | 
| 47694 | 552 | fix D assume "D \<in> ?D" then show "D \<subseteq> space M" | 
| 42982 | 553 | using sets_into_space by auto | 
| 554 | next | |
| 47694 | 555 | show "space M \<in> ?D" | 
| 42982 | 556 | using prob_space `X \<subseteq> space M` by (simp add: Int_absorb2) | 
| 557 | next | |
| 47694 | 558 | fix A assume A: "A \<in> ?D" | 
| 42982 | 559 | have "prob (X \<inter> (space M - A)) = prob (X - (X \<inter> A))" | 
| 560 | using `X \<subseteq> space M` by (auto intro!: arg_cong[where f=prob]) | |
| 561 | also have "\<dots> = prob X - prob (X \<inter> A)" | |
| 562 | using X_in A by (intro finite_measure_Diff) auto | |
| 563 | also have "\<dots> = prob X * prob (space M) - prob X * prob A" | |
| 564 | using A prob_space by auto | |
| 565 | also have "\<dots> = prob X * prob (space M - A)" | |
| 566 | using X_in A sets_into_space | |
| 567 | by (subst finite_measure_Diff) (auto simp: field_simps) | |
| 47694 | 568 | finally show "space M - A \<in> ?D" | 
| 42982 | 569 | using A `X \<subseteq> space M` by auto | 
| 570 | next | |
| 47694 | 571 | fix F :: "nat \<Rightarrow> 'a set" assume dis: "disjoint_family F" and "range F \<subseteq> ?D" | 
| 42982 | 572 | then have F: "range F \<subseteq> events" "\<And>i. prob (X \<inter> F i) = prob X * prob (F i)" | 
| 573 | by auto | |
| 574 | have "(\<lambda>i. prob (X \<inter> F i)) sums prob (\<Union>i. X \<inter> F i)" | |
| 575 | proof (rule finite_measure_UNION) | |
| 576 | show "range (\<lambda>i. X \<inter> F i) \<subseteq> events" | |
| 577 | using F X_in by auto | |
| 578 | show "disjoint_family (\<lambda>i. X \<inter> F i)" | |
| 579 | using dis by (rule disjoint_family_on_bisimulation) auto | |
| 580 | qed | |
| 581 | with F have "(\<lambda>i. prob X * prob (F i)) sums prob (X \<inter> (\<Union>i. F i))" | |
| 582 | by simp | |
| 583 | 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: 
43920diff
changeset | 584 | by (intro sums_mult finite_measure_UNION F dis) | 
| 42982 | 585 | ultimately have "prob (X \<inter> (\<Union>i. F i)) = prob X * prob (\<Union>i. F i)" | 
| 586 | by (auto dest!: sums_unique) | |
| 47694 | 587 | with F show "(\<Union>i. F i) \<in> ?D" | 
| 42982 | 588 | by auto | 
| 589 | qed | |
| 590 | ||
| 591 |   { fix n
 | |
| 592 |     have "indep_sets (\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) UNIV"
 | |
| 593 | proof (rule indep_sets_collect_sigma) | |
| 594 |       have *: "(\<Union>b. case b of True \<Rightarrow> {..n} | False \<Rightarrow> {Suc n..}) = UNIV" (is "?U = _")
 | |
| 595 | by (simp split: bool.split add: set_eq_iff) (metis not_less_eq_eq) | |
| 596 | with indep show "indep_sets A ?U" by simp | |
| 597 |       show "disjoint_family (bool_case {..n} {Suc n..})"
 | |
| 598 | unfolding disjoint_family_on_def by (auto split: bool.split) | |
| 599 | fix m | |
| 47694 | 600 | show "Int_stable (A m)" | 
| 42982 | 601 | unfolding Int_stable_def using A.Int by auto | 
| 602 | qed | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 603 |     also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) =
 | 
| 42982 | 604 |       bool_case (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
 | 
| 605 | by (auto intro!: ext split: bool.split) | |
| 606 |     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))"
 | |
| 607 | unfolding indep_set_def by simp | |
| 608 | ||
| 47694 | 609 |     have "sigma_sets (space M) (\<Union>m\<in>{..n}. A m) \<subseteq> ?D"
 | 
| 42982 | 610 | proof (simp add: subset_eq, rule) | 
| 611 |       fix D assume D: "D \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
 | |
| 612 |       have "X \<in> sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m)"
 | |
| 49772 | 613 | using X unfolding tail_events_def by simp | 
| 42982 | 614 | from indep_setD[OF indep D this] indep_setD_ev1[OF indep] D | 
| 615 | show "D \<in> events \<and> prob (X \<inter> D) = prob X * prob D" | |
| 616 | by (auto simp add: ac_simps) | |
| 617 | qed } | |
| 47694 | 618 |   then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> ?D" (is "?A \<subseteq> _")
 | 
| 42982 | 619 | by auto | 
| 620 | ||
| 49772 | 621 | note `X \<in> tail_events A` | 
| 47694 | 622 |   also {
 | 
| 623 |     have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> sigma_sets (space M) ?A"
 | |
| 624 | by (intro sigma_sets_subseteq UN_mono) auto | |
| 49772 | 625 | then have "tail_events A \<subseteq> sigma_sets (space M) ?A" | 
| 626 | unfolding tail_events_def by auto } | |
| 47694 | 627 | also have "sigma_sets (space M) ?A = dynkin (space M) ?A" | 
| 42982 | 628 | proof (rule sigma_eq_dynkin) | 
| 629 |     { fix B n assume "B \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
 | |
| 630 | then have "B \<subseteq> space M" | |
| 47694 | 631 | by induct (insert A sets_into_space[of _ M], auto) } | 
| 632 | then show "?A \<subseteq> Pow (space M)" by auto | |
| 633 | show "Int_stable ?A" | |
| 42982 | 634 | proof (rule Int_stableI) | 
| 635 | fix a assume "a \<in> ?A" then guess n .. note a = this | |
| 636 | fix b assume "b \<in> ?A" then guess m .. note b = this | |
| 47694 | 637 |       interpret Amn: sigma_algebra "space M" "sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
 | 
| 638 | using A sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto | |
| 42982 | 639 |       have "sigma_sets (space M) (\<Union>i\<in>{..n}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
 | 
| 640 | by (intro sigma_sets_subseteq UN_mono) auto | |
| 641 |       with a have "a \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
 | |
| 642 | moreover | |
| 643 |       have "sigma_sets (space M) (\<Union>i\<in>{..m}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
 | |
| 644 | by (intro sigma_sets_subseteq UN_mono) auto | |
| 645 |       with b have "b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
 | |
| 646 |       ultimately have "a \<inter> b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
 | |
| 47694 | 647 | using Amn.Int[of a b] by simp | 
| 42982 | 648 |       then show "a \<inter> b \<in> (\<Union>n. sigma_sets (space M) (\<Union>i\<in>{..n}. A i))" by auto
 | 
| 649 | qed | |
| 650 | qed | |
| 47694 | 651 | also have "dynkin (space M) ?A \<subseteq> ?D" | 
| 652 | using `?A \<subseteq> ?D` by (auto intro!: D.dynkin_subset) | |
| 653 | finally show ?thesis by auto | |
| 42982 | 654 | qed | 
| 655 | ||
| 42985 | 656 | lemma (in prob_space) borel_0_1_law: | 
| 657 | fixes F :: "nat \<Rightarrow> 'a set" | |
| 49781 | 658 | assumes F2: "indep_events F UNIV" | 
| 42985 | 659 |   shows "prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 0 \<or> prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 1"
 | 
| 660 | proof (rule kolmogorov_0_1_law[of "\<lambda>i. sigma_sets (space M) { F i }"])
 | |
| 49781 | 661 | have F1: "range F \<subseteq> events" | 
| 662 | using F2 by (simp add: indep_events_def subset_eq) | |
| 47694 | 663 |   { fix i show "sigma_algebra (space M) (sigma_sets (space M) {F i})"
 | 
| 49781 | 664 |       using sigma_algebra_sigma_sets[of "{F i}" "space M"] F1 sets_into_space
 | 
| 47694 | 665 | by auto } | 
| 42985 | 666 |   show "indep_sets (\<lambda>i. sigma_sets (space M) {F i}) UNIV"
 | 
| 47694 | 667 | proof (rule indep_sets_sigma) | 
| 42985 | 668 |     show "indep_sets (\<lambda>i. {F i}) UNIV"
 | 
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 669 | unfolding indep_events_def_alt[symmetric] by fact | 
| 47694 | 670 |     fix i show "Int_stable {F i}"
 | 
| 42985 | 671 | unfolding Int_stable_def by simp | 
| 672 | qed | |
| 46731 | 673 |   let ?Q = "\<lambda>n. \<Union>i\<in>{n..}. F i"
 | 
| 49772 | 674 |   show "(\<Inter>n. \<Union>m\<in>{n..}. F m) \<in> tail_events (\<lambda>i. sigma_sets (space M) {F i})"
 | 
| 675 | unfolding tail_events_def | |
| 42985 | 676 | proof | 
| 677 | fix j | |
| 47694 | 678 |     interpret S: sigma_algebra "space M" "sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
 | 
| 49781 | 679 | using order_trans[OF F1 space_closed] | 
| 47694 | 680 | by (intro sigma_algebra_sigma_sets) (simp add: sigma_sets_singleton subset_eq) | 
| 42985 | 681 |     have "(\<Inter>n. ?Q n) = (\<Inter>n\<in>{j..}. ?Q n)"
 | 
| 682 | by (intro decseq_SucI INT_decseq_offset UN_mono) auto | |
| 47694 | 683 |     also have "\<dots> \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
 | 
| 49781 | 684 | using order_trans[OF F1 space_closed] | 
| 42985 | 685 | by (safe intro!: S.countable_INT S.countable_UN) | 
| 47694 | 686 | (auto simp: sigma_sets_singleton intro!: sigma_sets.Basic bexI) | 
| 42985 | 687 |     finally show "(\<Inter>n. ?Q n) \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
 | 
| 47694 | 688 | by simp | 
| 42985 | 689 | qed | 
| 690 | qed | |
| 691 | ||
| 42987 | 692 | lemma (in prob_space) indep_sets_finite: | 
| 693 |   assumes I: "I \<noteq> {}" "finite I"
 | |
| 694 | and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> events" "\<And>i. i \<in> I \<Longrightarrow> space M \<in> F i" | |
| 695 | 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)))" | |
| 696 | proof | |
| 697 | assume *: "indep_sets F I" | |
| 698 | from I show "\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j))" | |
| 699 | by (intro indep_setsD[OF *] ballI) auto | |
| 700 | next | |
| 701 | assume indep: "\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j))" | |
| 702 | show "indep_sets F I" | |
| 703 | proof (rule indep_setsI[OF F(1)]) | |
| 704 |     fix A J assume J: "J \<noteq> {}" "J \<subseteq> I" "finite J"
 | |
| 705 | assume A: "\<forall>j\<in>J. A j \<in> F j" | |
| 46731 | 706 | let ?A = "\<lambda>j. if j \<in> J then A j else space M" | 
| 42987 | 707 | have "prob (\<Inter>j\<in>I. ?A j) = prob (\<Inter>j\<in>J. A j)" | 
| 708 | using subset_trans[OF F(1) space_closed] J A | |
| 709 | by (auto intro!: arg_cong[where f=prob] split: split_if_asm) blast | |
| 710 | also | |
| 711 | from A F have "(\<lambda>j. if j \<in> J then A j else space M) \<in> Pi I F" (is "?A \<in> _") | |
| 712 | by (auto split: split_if_asm) | |
| 713 | with indep have "prob (\<Inter>j\<in>I. ?A j) = (\<Prod>j\<in>I. prob (?A j))" | |
| 714 | by auto | |
| 715 | also have "\<dots> = (\<Prod>j\<in>J. prob (A j))" | |
| 716 | unfolding if_distrib setprod.If_cases[OF `finite I`] | |
| 717 | using prob_space `J \<subseteq> I` by (simp add: Int_absorb1 setprod_1) | |
| 718 | finally show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" .. | |
| 719 | qed | |
| 720 | qed | |
| 721 | ||
| 42989 | 722 | lemma (in prob_space) indep_vars_finite: | 
| 42987 | 723 | fixes I :: "'i set" | 
| 724 |   assumes I: "I \<noteq> {}" "finite I"
 | |
| 47694 | 725 | and M': "\<And>i. i \<in> I \<Longrightarrow> sets (M' i) = sigma_sets (space (M' i)) (E i)" | 
| 726 | and rv: "\<And>i. i \<in> I \<Longrightarrow> random_variable (M' i) (X i)" | |
| 727 | and Int_stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (E i)" | |
| 728 | 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))" | |
| 729 | shows "indep_vars M' X I \<longleftrightarrow> | |
| 730 | (\<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 | 731 | proof - | 
| 732 | from rv have X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> space M \<rightarrow> space (M' i)" | |
| 733 | unfolding measurable_def by simp | |
| 734 | ||
| 735 |   { fix i assume "i\<in>I"
 | |
| 47694 | 736 | from closed[OF `i \<in> I`] | 
| 737 |     have "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}
 | |
| 738 |       = sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> E i}"
 | |
| 739 | unfolding sigma_sets_vimage_commute[OF X, OF `i \<in> I`, symmetric] M'[OF `i \<in> I`] | |
| 42987 | 740 | by (subst sigma_sets_sigma_sets_eq) auto } | 
| 47694 | 741 | note sigma_sets_X = this | 
| 42987 | 742 | |
| 743 |   { fix i assume "i\<in>I"
 | |
| 47694 | 744 |     have "Int_stable {X i -` A \<inter> space M |A. A \<in> E i}"
 | 
| 42987 | 745 | proof (rule Int_stableI) | 
| 47694 | 746 |       fix a assume "a \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
 | 
| 747 | then obtain A where "a = X i -` A \<inter> space M" "A \<in> E i" by auto | |
| 42987 | 748 | moreover | 
| 47694 | 749 |       fix b assume "b \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
 | 
| 750 | then obtain B where "b = X i -` B \<inter> space M" "B \<in> E i" by auto | |
| 42987 | 751 | moreover | 
| 752 | have "(X i -` A \<inter> space M) \<inter> (X i -` B \<inter> space M) = X i -` (A \<inter> B) \<inter> space M" by auto | |
| 753 | moreover note Int_stable[OF `i \<in> I`] | |
| 754 | ultimately | |
| 47694 | 755 |       show "a \<inter> b \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
 | 
| 42987 | 756 | by (auto simp del: vimage_Int intro!: exI[of _ "A \<inter> B"] dest: Int_stableD) | 
| 757 | qed } | |
| 47694 | 758 | note indep_sets_X = indep_sets_sigma_sets_iff[OF this] | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 759 | |
| 42987 | 760 |   { fix i assume "i \<in> I"
 | 
| 47694 | 761 |     { fix A assume "A \<in> E i"
 | 
| 762 | with M'[OF `i \<in> I`] have "A \<in> sets (M' i)" by auto | |
| 42987 | 763 | moreover | 
| 47694 | 764 | from rv[OF `i\<in>I`] have "X i \<in> measurable M (M' i)" by auto | 
| 42987 | 765 | ultimately | 
| 766 | have "X i -` A \<inter> space M \<in> sets M" by (auto intro: measurable_sets) } | |
| 767 | with X[OF `i\<in>I`] space[OF `i\<in>I`] | |
| 47694 | 768 |     have "{X i -` A \<inter> space M |A. A \<in> E i} \<subseteq> events"
 | 
| 769 |       "space M \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
 | |
| 42987 | 770 | by (auto intro!: exI[of _ "space (M' i)"]) } | 
| 47694 | 771 | note indep_sets_finite_X = indep_sets_finite[OF I this] | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 772 | |
| 47694 | 773 |   have "(\<forall>A\<in>\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> E i}. prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))) =
 | 
| 774 | (\<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 | 775 | (is "?L = ?R") | 
| 776 | proof safe | |
| 47694 | 777 | fix A assume ?L and A: "A \<in> (\<Pi> i\<in>I. E i)" | 
| 42987 | 778 |     from `?L`[THEN bspec, of "\<lambda>i. X i -` A i \<inter> space M"] A `I \<noteq> {}`
 | 
| 779 | 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))" | |
| 780 | by (auto simp add: Pi_iff) | |
| 781 | next | |
| 47694 | 782 |     fix A assume ?R and A: "A \<in> (\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> E i})"
 | 
| 783 | 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 | 784 | from bchoice[OF this] obtain B where B: "\<forall>i\<in>I. A i = X i -` B i \<inter> space M" | 
| 47694 | 785 | "B \<in> (\<Pi> i\<in>I. E i)" by auto | 
| 42987 | 786 |     from `?R`[THEN bspec, OF B(2)] B(1) `I \<noteq> {}`
 | 
| 787 | show "prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))" | |
| 788 | by simp | |
| 789 | qed | |
| 790 |   then show ?thesis using `I \<noteq> {}`
 | |
| 47694 | 791 | by (simp add: rv indep_vars_def indep_sets_X sigma_sets_X indep_sets_finite_X cong: indep_sets_cong) | 
| 42988 | 792 | qed | 
| 793 | ||
| 42989 | 794 | lemma (in prob_space) indep_vars_compose: | 
| 795 | assumes "indep_vars M' X I" | |
| 47694 | 796 | assumes rv: "\<And>i. i \<in> I \<Longrightarrow> Y i \<in> measurable (M' i) (N i)" | 
| 42989 | 797 | shows "indep_vars N (\<lambda>i. Y i \<circ> X i) I" | 
| 798 | unfolding indep_vars_def | |
| 42988 | 799 | proof | 
| 42989 | 800 | from rv `indep_vars M' X I` | 
| 42988 | 801 | show "\<forall>i\<in>I. random_variable (N i) (Y i \<circ> X i)" | 
| 47694 | 802 | by (auto simp: indep_vars_def) | 
| 42988 | 803 | |
| 804 |   have "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
 | |
| 42989 | 805 | using `indep_vars M' X I` by (simp add: indep_vars_def) | 
| 42988 | 806 |   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"
 | 
| 807 | proof (rule indep_sets_mono_sets) | |
| 808 | fix i assume "i \<in> I" | |
| 42989 | 809 | with `indep_vars M' X I` have X: "X i \<in> space M \<rightarrow> space (M' i)" | 
| 810 | unfolding indep_vars_def measurable_def by auto | |
| 42988 | 811 |     { fix A assume "A \<in> sets (N i)"
 | 
| 812 | 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)" | |
| 813 | by (intro exI[of _ "Y i -` A \<inter> space (M' i)"]) | |
| 814 | (auto simp: vimage_compose intro!: measurable_sets rv `i \<in> I` funcset_mem[OF X]) } | |
| 815 |     then show "sigma_sets (space M) {(Y i \<circ> X i) -` A \<inter> space M |A. A \<in> sets (N i)} \<subseteq>
 | |
| 816 |       sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
 | |
| 817 | by (intro sigma_sets_subseteq) (auto simp: vimage_compose) | |
| 818 | qed | |
| 819 | qed | |
| 820 | ||
| 47694 | 821 | lemma (in prob_space) indep_varsD_finite: | 
| 42989 | 822 | assumes X: "indep_vars M' X I" | 
| 42988 | 823 |   assumes I: "I \<noteq> {}" "finite I" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M' i)"
 | 
| 824 | 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))" | |
| 825 | proof (rule indep_setsD) | |
| 826 |   show "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
 | |
| 42989 | 827 | using X by (auto simp: indep_vars_def) | 
| 42988 | 828 |   show "I \<subseteq> I" "I \<noteq> {}" "finite I" using I by auto
 | 
| 829 |   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 | 830 | using I by auto | 
| 42988 | 831 | qed | 
| 832 | ||
| 47694 | 833 | lemma (in prob_space) indep_varsD: | 
| 834 | assumes X: "indep_vars M' X I" | |
| 835 |   assumes I: "J \<noteq> {}" "finite J" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M' i)"
 | |
| 836 | 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))" | |
| 837 | proof (rule indep_setsD) | |
| 838 |   show "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
 | |
| 839 | using X by (auto simp: indep_vars_def) | |
| 840 |   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)}"
 | |
| 841 | using I by auto | |
| 842 | qed fact+ | |
| 843 | ||
| 844 | lemma prod_algebra_cong: | |
| 845 | assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))" | |
| 846 | shows "prod_algebra I M = prod_algebra J N" | |
| 847 | proof - | |
| 848 | have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)" | |
| 849 | using sets_eq_imp_space_eq[OF sets] by auto | |
| 850 | with sets show ?thesis unfolding `I = J` | |
| 851 | by (intro antisym prod_algebra_mono) auto | |
| 852 | qed | |
| 853 | ||
| 854 | lemma space_in_prod_algebra: | |
| 855 | "(\<Pi>\<^isub>E i\<in>I. space (M i)) \<in> prod_algebra I M" | |
| 856 | proof cases | |
| 857 |   assume "I = {}" then show ?thesis
 | |
| 858 | by (auto simp add: prod_algebra_def image_iff prod_emb_def) | |
| 859 | next | |
| 860 |   assume "I \<noteq> {}"
 | |
| 861 | then obtain i where "i \<in> I" by auto | |
| 862 |   then have "(\<Pi>\<^isub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i))"
 | |
| 863 | by (auto simp: prod_emb_def Pi_iff) | |
| 864 | also have "\<dots> \<in> prod_algebra I M" | |
| 865 | using `i \<in> I` by (intro prod_algebraI) auto | |
| 866 | finally show ?thesis . | |
| 867 | qed | |
| 868 | ||
| 869 | lemma (in prob_space) indep_vars_iff_distr_eq_PiM: | |
| 870 | fixes I :: "'i set" and X :: "'i \<Rightarrow> 'a \<Rightarrow> 'b" | |
| 871 |   assumes "I \<noteq> {}"
 | |
| 42988 | 872 | assumes rv: "\<And>i. random_variable (M' i) (X i)" | 
| 42989 | 873 | shows "indep_vars M' X I \<longleftrightarrow> | 
| 47694 | 874 | distr M (\<Pi>\<^isub>M i\<in>I. M' i) (\<lambda>x. \<lambda>i\<in>I. X i x) = (\<Pi>\<^isub>M i\<in>I. distr M (M' i) (X i))" | 
| 42988 | 875 | proof - | 
| 47694 | 876 | let ?P = "\<Pi>\<^isub>M i\<in>I. M' i" | 
| 877 | let ?X = "\<lambda>x. \<lambda>i\<in>I. X i x" | |
| 878 | let ?D = "distr M ?P ?X" | |
| 879 | have X: "random_variable ?P ?X" by (intro measurable_restrict rv) | |
| 880 | interpret D: prob_space ?D by (intro prob_space_distr X) | |
| 42988 | 881 | |
| 47694 | 882 | let ?D' = "\<lambda>i. distr M (M' i) (X i)" | 
| 883 | let ?P' = "\<Pi>\<^isub>M i\<in>I. distr M (M' i) (X i)" | |
| 884 | interpret D': prob_space "?D' i" for i by (intro prob_space_distr rv) | |
| 885 | interpret P: product_prob_space ?D' I .. | |
| 886 | ||
| 42988 | 887 | show ?thesis | 
| 47694 | 888 | proof | 
| 42989 | 889 | assume "indep_vars M' X I" | 
| 47694 | 890 | show "?D = ?P'" | 
| 891 | proof (rule measure_eqI_generator_eq) | |
| 892 | show "Int_stable (prod_algebra I M')" | |
| 893 | by (rule Int_stable_prod_algebra) | |
| 894 | show "prod_algebra I M' \<subseteq> Pow (space ?P)" | |
| 895 | using prod_algebra_sets_into_space by (simp add: space_PiM) | |
| 896 | show "sets ?D = sigma_sets (space ?P) (prod_algebra I M')" | |
| 897 | by (simp add: sets_PiM space_PiM) | |
| 898 | show "sets ?P' = sigma_sets (space ?P) (prod_algebra I M')" | |
| 899 | by (simp add: sets_PiM space_PiM cong: prod_algebra_cong) | |
| 900 | let ?A = "\<lambda>i. \<Pi>\<^isub>E i\<in>I. space (M' i)" | |
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 901 | show "range ?A \<subseteq> prod_algebra I M'" "(\<Union>i. ?A i) = space (Pi\<^isub>M I M')" | 
| 47694 | 902 | by (auto simp: space_PiM intro!: space_in_prod_algebra cong: prod_algebra_cong) | 
| 903 |       { fix i show "emeasure ?D (\<Pi>\<^isub>E i\<in>I. space (M' i)) \<noteq> \<infinity>" by auto }
 | |
| 904 | next | |
| 905 | fix E assume E: "E \<in> prod_algebra I M'" | |
| 906 | from prod_algebraE[OF E] guess J Y . note J = this | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 907 | |
| 47694 | 908 | from E have "E \<in> sets ?P" by (auto simp: sets_PiM) | 
| 909 | then have "emeasure ?D E = emeasure M (?X -` E \<inter> space M)" | |
| 910 | by (simp add: emeasure_distr X) | |
| 911 | also have "?X -` E \<inter> space M = (\<Inter>i\<in>J. X i -` Y i \<inter> space M)" | |
| 912 |         using J `I \<noteq> {}` measurable_space[OF rv] by (auto simp: prod_emb_def Pi_iff split: split_if_asm)
 | |
| 913 | 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))" | |
| 914 |         using `indep_vars M' X I` J `I \<noteq> {}` using indep_varsD[of M' X I J]
 | |
| 915 | by (auto simp: emeasure_eq_measure setprod_ereal) | |
| 916 | also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))" | |
| 917 | using rv J by (simp add: emeasure_distr) | |
| 918 | also have "\<dots> = emeasure ?P' E" | |
| 919 | using P.emeasure_PiM_emb[of J Y] J by (simp add: prod_emb_def) | |
| 920 | finally show "emeasure ?D E = emeasure ?P' E" . | |
| 42988 | 921 | qed | 
| 922 | next | |
| 47694 | 923 | assume "?D = ?P'" | 
| 924 | show "indep_vars M' X I" unfolding indep_vars_def | |
| 925 | proof (intro conjI indep_setsI ballI rv) | |
| 926 |       fix i show "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)} \<subseteq> events"
 | |
| 927 | by (auto intro!: sigma_sets_subset measurable_sets rv) | |
| 42988 | 928 | next | 
| 47694 | 929 |       fix J Y' assume J: "J \<noteq> {}" "J \<subseteq> I" "finite J"
 | 
| 930 |       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)}"
 | |
| 931 | have "\<forall>j\<in>J. \<exists>Y. Y' j = X j -` Y \<inter> space M \<and> Y \<in> sets (M' j)" | |
| 42988 | 932 | proof | 
| 47694 | 933 | fix j assume "j \<in> J" | 
| 934 | from Y'[rule_format, OF this] rv[of j] | |
| 935 | show "\<exists>Y. Y' j = X j -` Y \<inter> space M \<and> Y \<in> sets (M' j)" | |
| 936 | by (subst (asm) sigma_sets_vimage_commute[symmetric, of _ _ "space (M' j)"]) | |
| 937 | (auto dest: measurable_space simp: sigma_sets_eq) | |
| 42988 | 938 | qed | 
| 47694 | 939 | from bchoice[OF this] obtain Y where | 
| 940 | 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 | |
| 941 | let ?E = "prod_emb I M' J (Pi\<^isub>E J Y)" | |
| 942 | from Y have "(\<Inter>j\<in>J. Y' j) = ?X -` ?E \<inter> space M" | |
| 943 |         using J `I \<noteq> {}` measurable_space[OF rv] by (auto simp: prod_emb_def Pi_iff split: split_if_asm)
 | |
| 944 | then have "emeasure M (\<Inter>j\<in>J. Y' j) = emeasure M (?X -` ?E \<inter> space M)" | |
| 945 | by simp | |
| 946 | also have "\<dots> = emeasure ?D ?E" | |
| 947 | using Y J by (intro emeasure_distr[symmetric] X sets_PiM_I) auto | |
| 948 | also have "\<dots> = emeasure ?P' ?E" | |
| 949 | using `?D = ?P'` by simp | |
| 950 | also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))" | |
| 951 | using P.emeasure_PiM_emb[of J Y] J Y by (simp add: prod_emb_def) | |
| 952 | also have "\<dots> = (\<Prod> i\<in>J. emeasure M (Y' i))" | |
| 953 | using rv J Y by (simp add: emeasure_distr) | |
| 954 | finally have "emeasure M (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. emeasure M (Y' i))" . | |
| 955 | then show "prob (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. prob (Y' i))" | |
| 956 | by (auto simp: emeasure_eq_measure setprod_ereal) | |
| 42988 | 957 | qed | 
| 958 | qed | |
| 42987 | 959 | qed | 
| 960 | ||
| 42989 | 961 | lemma (in prob_space) indep_varD: | 
| 962 | assumes indep: "indep_var Ma A Mb B" | |
| 963 | assumes sets: "Xa \<in> sets Ma" "Xb \<in> sets Mb" | |
| 964 | shows "prob ((\<lambda>x. (A x, B x)) -` (Xa \<times> Xb) \<inter> space M) = | |
| 965 | prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)" | |
| 966 | proof - | |
| 967 | have "prob ((\<lambda>x. (A x, B x)) -` (Xa \<times> Xb) \<inter> space M) = | |
| 968 | prob (\<Inter>i\<in>UNIV. (bool_case A B i -` bool_case Xa Xb i \<inter> space M))" | |
| 969 | by (auto intro!: arg_cong[where f=prob] simp: UNIV_bool) | |
| 970 | also have "\<dots> = (\<Prod>i\<in>UNIV. prob (bool_case A B i -` bool_case Xa Xb i \<inter> space M))" | |
| 971 | using indep unfolding indep_var_def | |
| 972 | by (rule indep_varsD) (auto split: bool.split intro: sets) | |
| 973 | also have "\<dots> = prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)" | |
| 974 | unfolding UNIV_bool by simp | |
| 975 | finally show ?thesis . | |
| 976 | qed | |
| 977 | ||
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 978 | lemma (in prob_space) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 979 | assumes "indep_var S X T Y" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 980 | shows indep_var_rv1: "random_variable S X" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 981 | and indep_var_rv2: "random_variable T Y" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 982 | proof - | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 983 | have "\<forall>i\<in>UNIV. random_variable (bool_case S T i) (bool_case X Y i)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 984 | using assms unfolding indep_var_def indep_vars_def by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 985 | then show "random_variable S X" "random_variable T Y" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 986 | unfolding UNIV_bool by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 987 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 988 | |
| 47694 | 989 | lemma measurable_bool_case[simp, intro]: | 
| 990 | "(\<lambda>(x, y). bool_case x y) \<in> measurable (M \<Otimes>\<^isub>M N) (Pi\<^isub>M UNIV (bool_case M N))" | |
| 991 | (is "?f \<in> measurable ?B ?P") | |
| 992 | proof (rule measurable_PiM_single) | |
| 993 | show "?f \<in> space ?B \<rightarrow> (\<Pi>\<^isub>E i\<in>UNIV. space (bool_case M N i))" | |
| 994 | by (auto simp: space_pair_measure extensional_def split: bool.split) | |
| 995 | fix i A assume "A \<in> sets (case i of True \<Rightarrow> M | False \<Rightarrow> N)" | |
| 996 |   moreover then have "{\<omega> \<in> space (M \<Otimes>\<^isub>M N). prod_case bool_case \<omega> i \<in> A}
 | |
| 997 | = (case i of True \<Rightarrow> A \<times> space N | False \<Rightarrow> space M \<times> A)" | |
| 998 | by (auto simp: space_pair_measure split: bool.split dest!: sets_into_space) | |
| 999 |   ultimately show "{\<omega> \<in> space (M \<Otimes>\<^isub>M N). prod_case bool_case \<omega> i \<in> A} \<in> sets ?B"
 | |
| 1000 | by (auto split: bool.split) | |
| 1001 | qed | |
| 1002 | ||
| 1003 | lemma borel_measurable_indicator': | |
| 1004 | "A \<in> sets N \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> (\<lambda>x. indicator A (f x)) \<in> borel_measurable M" | |
| 1005 | using measurable_comp[OF _ borel_measurable_indicator, of f M N A] by (auto simp add: comp_def) | |
| 1006 | ||
| 1007 | lemma (in product_sigma_finite) distr_component: | |
| 1008 |   "distr (M i) (Pi\<^isub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^isub>M {i} M" (is "?D = ?P")
 | |
| 1009 | proof (intro measure_eqI[symmetric]) | |
| 1010 |   interpret I: finite_product_sigma_finite M "{i}" by default simp
 | |
| 1011 | ||
| 1012 |   have eq: "\<And>x. x \<in> extensional {i} \<Longrightarrow> (\<lambda>j\<in>{i}. x i) = x"
 | |
| 1013 | by (auto simp: extensional_def restrict_def) | |
| 1014 | ||
| 1015 | fix A assume A: "A \<in> sets ?P" | |
| 1016 | then have "emeasure ?P A = (\<integral>\<^isup>+x. indicator A x \<partial>?P)" | |
| 1017 | by simp | |
| 1018 |   also have "\<dots> = (\<integral>\<^isup>+x. indicator ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i)) x \<partial>M i)" 
 | |
| 1019 | apply (subst product_positive_integral_singleton[symmetric]) | |
| 1020 | apply (force intro!: measurable_restrict measurable_sets A) | |
| 1021 | apply (auto intro!: positive_integral_cong simp: space_PiM indicator_def simp: eq) | |
| 1022 | done | |
| 1023 |   also have "\<dots> = emeasure (M i) ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i))"
 | |
| 1024 | by (force intro!: measurable_restrict measurable_sets A positive_integral_indicator) | |
| 1025 | also have "\<dots> = emeasure ?D A" | |
| 1026 | using A by (auto intro!: emeasure_distr[symmetric] measurable_restrict) | |
| 1027 |   finally show "emeasure (Pi\<^isub>M {i} M) A = emeasure ?D A" .
 | |
| 1028 | qed simp | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1029 | |
| 47694 | 1030 | lemma pair_measure_eqI: | 
| 1031 | assumes "sigma_finite_measure M1" "sigma_finite_measure M2" | |
| 1032 | assumes sets: "sets (M1 \<Otimes>\<^isub>M M2) = sets M" | |
| 1033 | assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)" | |
| 1034 | shows "M1 \<Otimes>\<^isub>M M2 = M" | |
| 1035 | proof - | |
| 1036 | interpret M1: sigma_finite_measure M1 by fact | |
| 1037 | interpret M2: sigma_finite_measure M2 by fact | |
| 1038 | interpret pair_sigma_finite M1 M2 by default | |
| 1039 |   from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
 | |
| 1040 |   let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
 | |
| 1041 | let ?P = "M1 \<Otimes>\<^isub>M M2" | |
| 1042 | show ?thesis | |
| 1043 | proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) | |
| 1044 | show "?E \<subseteq> Pow (space ?P)" | |
| 1045 | using space_closed[of M1] space_closed[of M2] by (auto simp: space_pair_measure) | |
| 1046 | show "sets ?P = sigma_sets (space ?P) ?E" | |
| 1047 | by (simp add: sets_pair_measure space_pair_measure) | |
| 1048 | then show "sets M = sigma_sets (space ?P) ?E" | |
| 1049 | using sets[symmetric] by simp | |
| 1050 | next | |
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49781diff
changeset | 1051 | show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>" | 
| 47694 | 1052 | using F by (auto simp: space_pair_measure) | 
| 1053 | next | |
| 1054 | fix X assume "X \<in> ?E" | |
| 1055 | then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto | |
| 1056 | then have "emeasure ?P X = emeasure M1 A * emeasure M2 B" | |
| 49776 | 1057 | by (simp add: M2.emeasure_pair_measure_Times) | 
| 47694 | 1058 | also have "\<dots> = emeasure M (A \<times> B)" | 
| 1059 | using A B emeasure by auto | |
| 1060 | finally show "emeasure ?P X = emeasure M X" | |
| 1061 | by simp | |
| 1062 | qed | |
| 1063 | qed | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1064 | |
| 47694 | 1065 | lemma pair_measure_eq_distr_PiM: | 
| 1066 | fixes M1 :: "'a measure" and M2 :: "'a measure" | |
| 1067 | assumes "sigma_finite_measure M1" "sigma_finite_measure M2" | |
| 1068 | shows "(M1 \<Otimes>\<^isub>M M2) = distr (Pi\<^isub>M UNIV (bool_case M1 M2)) (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. (x True, x False))" | |
| 1069 | (is "?P = ?D") | |
| 1070 | proof (rule pair_measure_eqI[OF assms]) | |
| 1071 | interpret B: product_sigma_finite "bool_case M1 M2" | |
| 1072 | unfolding product_sigma_finite_def using assms by (auto split: bool.split) | |
| 1073 | let ?B = "Pi\<^isub>M UNIV (bool_case M1 M2)" | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1074 | |
| 47694 | 1075 | have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)" | 
| 1076 | by auto | |
| 1077 | fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2" | |
| 1078 | have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (bool_case M1 M2 i) (bool_case A B i))" | |
| 1079 | by (simp add: UNIV_bool ac_simps) | |
| 1080 | also have "\<dots> = emeasure ?B (Pi\<^isub>E UNIV (bool_case A B))" | |
| 1081 | using A B by (subst B.emeasure_PiM) (auto split: bool.split) | |
| 1082 | also have "Pi\<^isub>E UNIV (bool_case A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B" | |
| 1083 | using A[THEN sets_into_space] B[THEN sets_into_space] | |
| 1084 | by (auto simp: Pi_iff all_bool_eq space_PiM split: bool.split) | |
| 1085 | finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)" | |
| 1086 | using A B | |
| 1087 | measurable_component_singleton[of True UNIV "bool_case M1 M2"] | |
| 1088 | measurable_component_singleton[of False UNIV "bool_case M1 M2"] | |
| 1089 | by (subst emeasure_distr) (auto simp: measurable_pair_iff) | |
| 1090 | qed simp | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1091 | |
| 47694 | 1092 | lemma measurable_Pair: | 
| 1093 | assumes rvs: "X \<in> measurable M S" "Y \<in> measurable M T" | |
| 1094 | shows "(\<lambda>x. (X x, Y x)) \<in> measurable M (S \<Otimes>\<^isub>M T)" | |
| 1095 | proof - | |
| 1096 | have [simp]: "fst \<circ> (\<lambda>x. (X x, Y x)) = (\<lambda>x. X x)" "snd \<circ> (\<lambda>x. (X x, Y x)) = (\<lambda>x. Y x)" | |
| 1097 | by auto | |
| 1098 | show " (\<lambda>x. (X x, Y x)) \<in> measurable M (S \<Otimes>\<^isub>M T)" | |
| 1099 | by (auto simp: measurable_pair_iff rvs) | |
| 1100 | qed | |
| 1101 | ||
| 1102 | lemma (in prob_space) indep_var_distribution_eq: | |
| 1103 | "indep_var S X T Y \<longleftrightarrow> random_variable S X \<and> random_variable T Y \<and> | |
| 1104 | distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))" (is "_ \<longleftrightarrow> _ \<and> _ \<and> ?S \<Otimes>\<^isub>M ?T = ?J") | |
| 1105 | proof safe | |
| 1106 | assume "indep_var S X T Y" | |
| 1107 | then show rvs: "random_variable S X" "random_variable T Y" | |
| 1108 | by (blast dest: indep_var_rv1 indep_var_rv2)+ | |
| 1109 | then have XY: "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))" | |
| 1110 | by (rule measurable_Pair) | |
| 1111 | ||
| 1112 | interpret X: prob_space ?S by (rule prob_space_distr) fact | |
| 1113 | interpret Y: prob_space ?T by (rule prob_space_distr) fact | |
| 1114 | interpret XY: pair_prob_space ?S ?T .. | |
| 1115 | show "?S \<Otimes>\<^isub>M ?T = ?J" | |
| 1116 | proof (rule pair_measure_eqI) | |
| 1117 | show "sigma_finite_measure ?S" .. | |
| 1118 | show "sigma_finite_measure ?T" .. | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1119 | |
| 47694 | 1120 | fix A B assume A: "A \<in> sets ?S" and B: "B \<in> sets ?T" | 
| 1121 | have "emeasure ?J (A \<times> B) = emeasure M ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)" | |
| 1122 | using A B by (intro emeasure_distr[OF XY]) auto | |
| 1123 | also have "\<dots> = emeasure M (X -` A \<inter> space M) * emeasure M (Y -` B \<inter> space M)" | |
| 1124 | using indep_varD[OF `indep_var S X T Y`, of A B] A B by (simp add: emeasure_eq_measure) | |
| 1125 | also have "\<dots> = emeasure ?S A * emeasure ?T B" | |
| 1126 | using rvs A B by (simp add: emeasure_distr) | |
| 1127 | finally show "emeasure ?S A * emeasure ?T B = emeasure ?J (A \<times> B)" by simp | |
| 1128 | qed simp | |
| 1129 | next | |
| 1130 | assume rvs: "random_variable S X" "random_variable T Y" | |
| 1131 | then have XY: "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))" | |
| 1132 | by (rule measurable_Pair) | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1133 | |
| 47694 | 1134 | let ?S = "distr M S X" and ?T = "distr M T Y" | 
| 1135 | interpret X: prob_space ?S by (rule prob_space_distr) fact | |
| 1136 | interpret Y: prob_space ?T by (rule prob_space_distr) fact | |
| 1137 | interpret XY: pair_prob_space ?S ?T .. | |
| 1138 | ||
| 1139 | assume "?S \<Otimes>\<^isub>M ?T = ?J" | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1140 | |
| 47694 | 1141 |   { fix S and X
 | 
| 1142 |     have "Int_stable {X -` A \<inter> space M |A. A \<in> sets S}"
 | |
| 1143 | proof (safe intro!: Int_stableI) | |
| 1144 | fix A B assume "A \<in> sets S" "B \<in> sets S" | |
| 1145 | 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" | |
| 1146 | by (intro exI[of _ "A \<inter> B"]) auto | |
| 1147 | qed } | |
| 1148 | note Int_stable = this | |
| 1149 | ||
| 1150 | show "indep_var S X T Y" unfolding indep_var_eq | |
| 1151 | proof (intro conjI indep_set_sigma_sets Int_stable rvs) | |
| 1152 |     show "indep_set {X -` A \<inter> space M |A. A \<in> sets S} {Y -` A \<inter> space M |A. A \<in> sets T}"
 | |
| 1153 | proof (safe intro!: indep_setI) | |
| 1154 |       { fix A assume "A \<in> sets S" then show "X -` A \<inter> space M \<in> sets M"
 | |
| 1155 | using `X \<in> measurable M S` by (auto intro: measurable_sets) } | |
| 1156 |       { fix A assume "A \<in> sets T" then show "Y -` A \<inter> space M \<in> sets M"
 | |
| 1157 | using `Y \<in> measurable M T` by (auto intro: measurable_sets) } | |
| 1158 | next | |
| 1159 | fix A B assume ab: "A \<in> sets S" "B \<in> sets T" | |
| 1160 | then have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) = emeasure ?J (A \<times> B)" | |
| 1161 | using XY by (auto simp add: emeasure_distr emeasure_eq_measure intro!: arg_cong[where f="prob"]) | |
| 1162 | also have "\<dots> = emeasure (?S \<Otimes>\<^isub>M ?T) (A \<times> B)" | |
| 1163 | unfolding `?S \<Otimes>\<^isub>M ?T = ?J` .. | |
| 1164 | also have "\<dots> = emeasure ?S A * emeasure ?T B" | |
| 49776 | 1165 | using ab by (simp add: Y.emeasure_pair_measure_Times) | 
| 47694 | 1166 | finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) = | 
| 1167 | prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)" | |
| 1168 | using rvs ab by (simp add: emeasure_eq_measure emeasure_distr) | |
| 1169 | qed | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1170 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42989diff
changeset | 1171 | qed | 
| 42989 | 1172 | |
| 42861 
16375b493b64
Add formalization of probabilistic independence for families of sets
 hoelzl parents: diff
changeset | 1173 | end |