| author | wenzelm | 
| Sun, 04 Sep 2011 19:36:19 +0200 | |
| changeset 44706 | fe319b45315c | 
| parent 43920 | cedb5cb948fd | 
| child 44890 | 22f665a2e91c | 
| permissions | -rw-r--r-- | 
| 42067 | 1 | (* Title: HOL/Probability/Information.thy | 
| 2 | Author: Johannes Hölzl, TU München | |
| 3 | Author: Armin Heller, TU München | |
| 4 | *) | |
| 5 | ||
| 6 | header {*Information theory*}
 | |
| 7 | ||
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 8 | theory Information | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
41095diff
changeset | 9 | imports | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 10 | Independent_Family | 
| 43556 
0d78c8d31d0d
move conditional expectation to its own theory file
 hoelzl parents: 
43340diff
changeset | 11 | Radon_Nikodym | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
41095diff
changeset | 12 | "~~/src/HOL/Library/Convex" | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 13 | begin | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 14 | |
| 39097 | 15 | lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y" | 
| 16 | by (subst log_le_cancel_iff) auto | |
| 17 | ||
| 18 | lemma log_less: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x < y \<Longrightarrow> log a x < log a y" | |
| 19 | by (subst log_less_cancel_iff) auto | |
| 20 | ||
| 21 | lemma setsum_cartesian_product': | |
| 22 | "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)" | |
| 23 | unfolding setsum_cartesian_product by simp | |
| 24 | ||
| 36624 | 25 | section "Convex theory" | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 26 | |
| 36624 | 27 | lemma log_setsum: | 
| 28 |   assumes "finite s" "s \<noteq> {}"
 | |
| 29 | assumes "b > 1" | |
| 30 | assumes "(\<Sum> i \<in> s. a i) = 1" | |
| 31 | assumes "\<And> i. i \<in> s \<Longrightarrow> a i \<ge> 0" | |
| 32 |   assumes "\<And> i. i \<in> s \<Longrightarrow> y i \<in> {0 <..}"
 | |
| 33 | shows "log b (\<Sum> i \<in> s. a i * y i) \<ge> (\<Sum> i \<in> s. a i * log b (y i))" | |
| 34 | proof - | |
| 35 |   have "convex_on {0 <..} (\<lambda> x. - log b x)"
 | |
| 36 | by (rule minus_log_convex[OF `b > 1`]) | |
| 37 | hence "- log b (\<Sum> i \<in> s. a i * y i) \<le> (\<Sum> i \<in> s. a i * - log b (y i))" | |
| 38 | using convex_on_setsum[of _ _ "\<lambda> x. - log b x"] assms pos_is_convex by fastsimp | |
| 39 | thus ?thesis by (auto simp add:setsum_negf le_imp_neg_le) | |
| 40 | qed | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 41 | |
| 36624 | 42 | lemma log_setsum': | 
| 43 |   assumes "finite s" "s \<noteq> {}"
 | |
| 44 | assumes "b > 1" | |
| 45 | assumes "(\<Sum> i \<in> s. a i) = 1" | |
| 46 | assumes pos: "\<And> i. i \<in> s \<Longrightarrow> 0 \<le> a i" | |
| 47 | "\<And> i. \<lbrakk> i \<in> s ; 0 < a i \<rbrakk> \<Longrightarrow> 0 < y i" | |
| 48 | shows "log b (\<Sum> i \<in> s. a i * y i) \<ge> (\<Sum> i \<in> s. a i * log b (y i))" | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 49 | proof - | 
| 36624 | 50 |   have "\<And>y. (\<Sum> i \<in> s - {i. a i = 0}. a i * y i) = (\<Sum> i \<in> s. a i * y i)"
 | 
| 51 | using assms by (auto intro!: setsum_mono_zero_cong_left) | |
| 52 |   moreover have "log b (\<Sum> i \<in> s - {i. a i = 0}. a i * y i) \<ge> (\<Sum> i \<in> s - {i. a i = 0}. a i * log b (y i))"
 | |
| 53 | proof (rule log_setsum) | |
| 54 |     have "setsum a (s - {i. a i = 0}) = setsum a s"
 | |
| 55 | using assms(1) by (rule setsum_mono_zero_cong_left) auto | |
| 56 |     thus sum_1: "setsum a (s - {i. a i = 0}) = 1"
 | |
| 57 |       "finite (s - {i. a i = 0})" using assms by simp_all
 | |
| 58 | ||
| 59 |     show "s - {i. a i = 0} \<noteq> {}"
 | |
| 60 | proof | |
| 61 |       assume *: "s - {i. a i = 0} = {}"
 | |
| 62 |       hence "setsum a (s - {i. a i = 0}) = 0" by (simp add: * setsum_empty)
 | |
| 63 | with sum_1 show False by simp | |
| 38656 | 64 | qed | 
| 36624 | 65 | |
| 66 |     fix i assume "i \<in> s - {i. a i = 0}"
 | |
| 67 | hence "i \<in> s" "a i \<noteq> 0" by simp_all | |
| 68 |     thus "0 \<le> a i" "y i \<in> {0<..}" using pos[of i] by auto
 | |
| 69 | qed fact+ | |
| 70 | ultimately show ?thesis by simp | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 71 | qed | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 72 | |
| 36624 | 73 | lemma log_setsum_divide: | 
| 74 |   assumes "finite S" and "S \<noteq> {}" and "1 < b"
 | |
| 75 | assumes "(\<Sum>x\<in>S. g x) = 1" | |
| 76 | assumes pos: "\<And>x. x \<in> S \<Longrightarrow> g x \<ge> 0" "\<And>x. x \<in> S \<Longrightarrow> f x \<ge> 0" | |
| 77 | assumes g_pos: "\<And>x. \<lbrakk> x \<in> S ; 0 < g x \<rbrakk> \<Longrightarrow> 0 < f x" | |
| 78 | shows "- (\<Sum>x\<in>S. g x * log b (g x / f x)) \<le> log b (\<Sum>x\<in>S. f x)" | |
| 79 | proof - | |
| 80 | have log_mono: "\<And>x y. 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log b x \<le> log b y" | |
| 81 | using `1 < b` by (subst log_le_cancel_iff) auto | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 82 | |
| 36624 | 83 | have "- (\<Sum>x\<in>S. g x * log b (g x / f x)) = (\<Sum>x\<in>S. g x * log b (f x / g x))" | 
| 84 | proof (unfold setsum_negf[symmetric], rule setsum_cong) | |
| 85 | fix x assume x: "x \<in> S" | |
| 86 | show "- (g x * log b (g x / f x)) = g x * log b (f x / g x)" | |
| 87 | proof (cases "g x = 0") | |
| 88 | case False | |
| 89 | with pos[OF x] g_pos[OF x] have "0 < f x" "0 < g x" by simp_all | |
| 90 | thus ?thesis using `1 < b` by (simp add: log_divide field_simps) | |
| 91 | qed simp | |
| 92 | qed rule | |
| 93 | also have "... \<le> log b (\<Sum>x\<in>S. g x * (f x / g x))" | |
| 94 | proof (rule log_setsum') | |
| 95 | fix x assume x: "x \<in> S" "0 < g x" | |
| 96 | with g_pos[OF x] show "0 < f x / g x" by (safe intro!: divide_pos_pos) | |
| 97 | qed fact+ | |
| 98 |   also have "... = log b (\<Sum>x\<in>S - {x. g x = 0}. f x)" using `finite S`
 | |
| 99 | by (auto intro!: setsum_mono_zero_cong_right arg_cong[where f="log b"] | |
| 100 | split: split_if_asm) | |
| 101 | also have "... \<le> log b (\<Sum>x\<in>S. f x)" | |
| 102 | proof (rule log_mono) | |
| 103 |     have "0 = (\<Sum>x\<in>S - {x. g x = 0}. 0)" by simp
 | |
| 104 |     also have "... < (\<Sum>x\<in>S - {x. g x = 0}. f x)" (is "_ < ?sum")
 | |
| 105 | proof (rule setsum_strict_mono) | |
| 106 |       show "finite (S - {x. g x = 0})" using `finite S` by simp
 | |
| 107 |       show "S - {x. g x = 0} \<noteq> {}"
 | |
| 108 | proof | |
| 109 |         assume "S - {x. g x = 0} = {}"
 | |
| 110 | hence "(\<Sum>x\<in>S. g x) = 0" by (subst setsum_0') auto | |
| 111 | with `(\<Sum>x\<in>S. g x) = 1` show False by simp | |
| 112 | qed | |
| 113 |       fix x assume "x \<in> S - {x. g x = 0}"
 | |
| 114 | thus "0 < f x" using g_pos[of x] pos(1)[of x] by auto | |
| 115 | qed | |
| 116 | finally show "0 < ?sum" . | |
| 117 |     show "(\<Sum>x\<in>S - {x. g x = 0}. f x) \<le> (\<Sum>x\<in>S. f x)"
 | |
| 118 | using `finite S` pos by (auto intro!: setsum_mono2) | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 119 | qed | 
| 36624 | 120 | finally show ?thesis . | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 121 | qed | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 122 | |
| 39097 | 123 | lemma split_pairs: | 
| 40859 | 124 | "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and | 
| 125 | "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto | |
| 38656 | 126 | |
| 127 | section "Information theory" | |
| 128 | ||
| 40859 | 129 | locale information_space = prob_space + | 
| 38656 | 130 | fixes b :: real assumes b_gt_1: "1 < b" | 
| 131 | ||
| 40859 | 132 | context information_space | 
| 38656 | 133 | begin | 
| 134 | ||
| 40859 | 135 | text {* Introduce some simplification rules for logarithm of base @{term b}. *}
 | 
| 136 | ||
| 137 | lemma log_neg_const: | |
| 138 | assumes "x \<le> 0" | |
| 139 | shows "log b x = log b 0" | |
| 36624 | 140 | proof - | 
| 40859 | 141 |   { fix u :: real
 | 
| 142 | have "x \<le> 0" by fact | |
| 143 | also have "0 < exp u" | |
| 144 | using exp_gt_zero . | |
| 145 | finally have "exp u \<noteq> x" | |
| 146 | by auto } | |
| 147 | then show "log b x = log b 0" | |
| 148 | by (simp add: log_def ln_def) | |
| 38656 | 149 | qed | 
| 150 | ||
| 40859 | 151 | lemma log_mult_eq: | 
| 152 | "log b (A * B) = (if 0 < A * B then log b \<bar>A\<bar> + log b \<bar>B\<bar> else log b 0)" | |
| 153 | using log_mult[of b "\<bar>A\<bar>" "\<bar>B\<bar>"] b_gt_1 log_neg_const[of "A * B"] | |
| 154 | by (auto simp: zero_less_mult_iff mult_le_0_iff) | |
| 38656 | 155 | |
| 40859 | 156 | lemma log_inverse_eq: | 
| 157 | "log b (inverse B) = (if 0 < B then - log b B else log b 0)" | |
| 158 | using log_inverse[of b B] log_neg_const[of "inverse B"] b_gt_1 by simp | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 159 | |
| 40859 | 160 | lemma log_divide_eq: | 
| 161 | "log b (A / B) = (if 0 < A * B then log b \<bar>A\<bar> - log b \<bar>B\<bar> else log b 0)" | |
| 162 | unfolding divide_inverse log_mult_eq log_inverse_eq abs_inverse | |
| 163 | by (auto simp: zero_less_mult_iff mult_le_0_iff) | |
| 38656 | 164 | |
| 40859 | 165 | lemmas log_simps = log_mult_eq log_inverse_eq log_divide_eq | 
| 38656 | 166 | |
| 167 | end | |
| 168 | ||
| 39097 | 169 | subsection "Kullback$-$Leibler divergence" | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 170 | |
| 39097 | 171 | text {* The Kullback$-$Leibler divergence is also known as relative entropy or
 | 
| 172 | Kullback$-$Leibler distance. *} | |
| 173 | ||
| 174 | definition | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 175 | "entropy_density b M \<nu> = log b \<circ> real \<circ> RN_deriv M \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 176 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 177 | definition | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 178 | "KL_divergence b M \<nu> = integral\<^isup>L (M\<lparr>measure := \<nu>\<rparr>) (entropy_density b M \<nu>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 179 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 180 | lemma (in information_space) measurable_entropy_density: | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 181 | assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 182 | assumes ac: "absolutely_continuous \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 183 | shows "entropy_density b M \<nu> \<in> borel_measurable M" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 184 | proof - | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 185 | interpret \<nu>: prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 186 | have "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by fact | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 187 | from RN_deriv[OF this ac] b_gt_1 show ?thesis | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 188 | unfolding entropy_density_def | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 189 | by (intro measurable_comp) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 190 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 191 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 192 | lemma (in information_space) KL_gt_0: | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 193 | assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 194 | assumes ac: "absolutely_continuous \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 195 | assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 196 | assumes A: "A \<in> sets M" "\<nu> A \<noteq> \<mu> A" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 197 | shows "0 < KL_divergence b M \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 198 | proof - | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 199 | interpret \<nu>: prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 200 | have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 201 | have fms: "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" by default | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 202 | note RN = RN_deriv[OF ms ac] | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 203 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 204 | from real_RN_deriv[OF fms ac] guess D . note D = this | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 205 | with absolutely_continuous_AE[OF ms] ac | 
| 43920 | 206 | have D\<nu>: "AE x in M\<lparr>measure := \<nu>\<rparr>. RN_deriv M \<nu> x = ereal (D x)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 207 | by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 208 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 209 | def f \<equiv> "\<lambda>x. if D x = 0 then 1 else 1 / D x" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 210 | with D have f_borel: "f \<in> borel_measurable M" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 211 | by (auto intro!: measurable_If) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 212 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 213 | have "KL_divergence b M \<nu> = 1 / ln b * (\<integral> x. ln b * entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 214 | unfolding KL_divergence_def using int b_gt_1 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 215 | by (simp add: integral_cmult) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 216 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 217 |   { fix A assume "A \<in> sets M"
 | 
| 43920 | 218 | with RN D have "\<nu>.\<mu> A = (\<integral>\<^isup>+ x. ereal (D x) * indicator A x \<partial>M)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 219 | by (auto intro!: positive_integral_cong_AE) } | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 220 | note D_density = this | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 221 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 222 | have ln_entropy: "(\<lambda>x. ln b * entropy_density b M \<nu> x) \<in> borel_measurable M" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 223 | using measurable_entropy_density[OF ps ac] by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 224 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 225 | have "integrable (M\<lparr>measure := \<nu>\<rparr>) (\<lambda>x. ln b * entropy_density b M \<nu> x)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 226 | using int by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 227 | moreover have "integrable (M\<lparr>measure := \<nu>\<rparr>) (\<lambda>x. ln b * entropy_density b M \<nu> x) \<longleftrightarrow> | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 228 | integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 229 | using D D_density ln_entropy | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 230 | by (intro integral_translated_density) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 231 | ultimately have M_int: "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 232 | by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 233 | |
| 43920 | 234 | have D_neg: "(\<integral>\<^isup>+ x. ereal (- D x) \<partial>M) = 0" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 235 | using D by (subst positive_integral_0_iff_AE) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 236 | |
| 43920 | 237 | have "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = \<nu> (space M)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 238 | using RN D by (auto intro!: positive_integral_cong_AE) | 
| 43920 | 239 | then have D_pos: "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = 1" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 240 | using \<nu>.measure_space_1 by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 241 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 242 | have "integrable M D" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 243 | using D_pos D_neg D by (auto simp: integrable_def) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 244 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 245 | have "integral\<^isup>L M D = 1" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 246 | using D_pos D_neg by (auto simp: lebesgue_integral_def) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 247 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 248 |   let ?D_set = "{x\<in>space M. D x \<noteq> 0}"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 249 | have [simp, intro]: "?D_set \<in> sets M" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 250 | using D by (auto intro: sets_Collect) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 251 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 252 | have "0 \<le> 1 - \<mu>' ?D_set" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 253 | using prob_le_1 by (auto simp: field_simps) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 254 | also have "\<dots> = (\<integral> x. D x - indicator ?D_set x \<partial>M)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 255 | using `integrable M D` `integral\<^isup>L M D = 1` | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 256 | by (simp add: \<mu>'_def) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 257 | also have "\<dots> < (\<integral> x. D x * (ln b * entropy_density b M \<nu> x) \<partial>M)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 258 | proof (rule integral_less_AE) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 259 | show "integrable M (\<lambda>x. D x - indicator ?D_set x)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 260 | using `integrable M D` | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 261 | by (intro integral_diff integral_indicator) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 262 | next | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 263 | show "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 264 | by fact | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 265 | next | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 266 |     show "\<mu> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<noteq> 0"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 267 | proof | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 268 |       assume eq_0: "\<mu> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} = 0"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 269 | then have disj: "AE x. D x = 1 \<or> D x = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 270 | using D(1) by (auto intro!: AE_I[OF subset_refl] sets_Collect) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 271 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 272 |       have "\<mu> {x\<in>space M. D x = 1} = (\<integral>\<^isup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 273 | using D(1) by auto | 
| 43920 | 274 |       also have "\<dots> = (\<integral>\<^isup>+ x. ereal (D x) * indicator {x\<in>space M. D x \<noteq> 0} x \<partial>M)"
 | 
| 275 | using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_ereal_def) | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 276 |       also have "\<dots> = \<nu> {x\<in>space M. D x \<noteq> 0}"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 277 | using D(1) D_density by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 278 | also have "\<dots> = \<nu> (space M)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 279 | using D_density D(1) by (auto intro!: positive_integral_cong simp: indicator_def) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 280 | finally have "AE x. D x = 1" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 281 | using D(1) \<nu>.measure_space_1 by (intro AE_I_eq_1) auto | 
| 43920 | 282 | then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. ereal (D x) * indicator A x\<partial>M)" | 
| 283 | by (intro positive_integral_cong_AE) (auto simp: one_ereal_def[symmetric]) | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 284 | also have "\<dots> = \<nu> A" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 285 | using `A \<in> sets M` D_density by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 286 | finally show False using `A \<in> sets M` `\<nu> A \<noteq> \<mu> A` by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 287 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 288 |     show "{x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<in> sets M"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 289 | using D(1) by (auto intro: sets_Collect) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 290 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 291 |     show "AE t. t \<in> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<longrightarrow>
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 292 | D t - indicator ?D_set t \<noteq> D t * (ln b * entropy_density b M \<nu> t)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 293 | using D(2) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 294 | proof (elim AE_mp, safe intro!: AE_I2) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 295 | fix t assume Dt: "t \<in> space M" "D t \<noteq> 1" "D t \<noteq> 0" | 
| 43920 | 296 | and RN: "RN_deriv M \<nu> t = ereal (D t)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 297 | and eq: "D t - indicator ?D_set t = D t * (ln b * entropy_density b M \<nu> t)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 298 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 299 | have "D t - 1 = D t - indicator ?D_set t" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 300 | using Dt by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 301 | also note eq | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 302 | also have "D t * (ln b * entropy_density b M \<nu> t) = - D t * ln (1 / D t)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 303 | using RN b_gt_1 `D t \<noteq> 0` `0 \<le> D t` | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 304 | by (simp add: entropy_density_def log_def ln_div less_le) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 305 | finally have "ln (1 / D t) = 1 / D t - 1" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 306 | using `D t \<noteq> 0` by (auto simp: field_simps) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 307 | from ln_eq_minus_one[OF _ this] `D t \<noteq> 0` `0 \<le> D t` `D t \<noteq> 1` | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 308 | show False by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 309 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 310 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 311 | show "AE t. D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 312 | using D(2) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 313 | proof (elim AE_mp, intro AE_I2 impI) | 
| 43920 | 314 | fix t assume "t \<in> space M" and RN: "RN_deriv M \<nu> t = ereal (D t)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 315 | show "D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 316 | proof cases | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 317 | assume asm: "D t \<noteq> 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 318 | then have "0 < D t" using `0 \<le> D t` by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 319 | then have "0 < 1 / D t" by auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 320 | have "D t - indicator ?D_set t \<le> - D t * (1 / D t - 1)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 321 | using asm `t \<in> space M` by (simp add: field_simps) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 322 | also have "- D t * (1 / D t - 1) \<le> - D t * ln (1 / D t)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 323 | using ln_le_minus_one `0 < 1 / D t` by (intro mult_left_mono_neg) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 324 | also have "\<dots> = D t * (ln b * entropy_density b M \<nu> t)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 325 | using `0 < D t` RN b_gt_1 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 326 | by (simp_all add: log_def ln_div entropy_density_def) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 327 | finally show ?thesis by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 328 | qed simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 329 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 330 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 331 | also have "\<dots> = (\<integral> x. ln b * entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 332 | using D D_density ln_entropy | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 333 | by (intro integral_translated_density[symmetric]) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 334 | also have "\<dots> = ln b * (\<integral> x. entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 335 | using int by (rule \<nu>.integral_cmult) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 336 | finally show "0 < KL_divergence b M \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 337 | using b_gt_1 by (auto simp: KL_divergence_def zero_less_mult_iff) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 338 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 339 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 340 | lemma (in sigma_finite_measure) KL_eq_0: | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 341 | assumes eq: "\<forall>A\<in>sets M. \<nu> A = measure M A" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 342 | shows "KL_divergence b M \<nu> = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 343 | proof - | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 344 | have "AE x. 1 = RN_deriv M \<nu> x" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 345 | proof (rule RN_deriv_unique) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 346 | show "measure_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 347 | using eq by (intro measure_space_cong) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 348 | show "absolutely_continuous \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 349 | unfolding absolutely_continuous_def using eq by auto | 
| 43920 | 350 | show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x. 0 \<le> (1 :: ereal)" by auto | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 351 | fix A assume "A \<in> sets M" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 352 | with eq show "\<nu> A = \<integral>\<^isup>+ x. 1 * indicator A x \<partial>M" by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 353 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 354 | then have "AE x. log b (real (RN_deriv M \<nu> x)) = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 355 | by (elim AE_mp) simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 356 | from integral_cong_AE[OF this] | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 357 | have "integral\<^isup>L M (entropy_density b M \<nu>) = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 358 | by (simp add: entropy_density_def comp_def) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 359 | with eq show "KL_divergence b M \<nu> = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 360 | unfolding KL_divergence_def | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 361 | by (subst integral_cong_measure) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 362 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 363 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 364 | lemma (in information_space) KL_eq_0_imp: | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 365 | assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 366 | assumes ac: "absolutely_continuous \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 367 | assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 368 | assumes KL: "KL_divergence b M \<nu> = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 369 | shows "\<forall>A\<in>sets M. \<nu> A = \<mu> A" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 370 | by (metis less_imp_neq KL_gt_0 assms) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 371 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 372 | lemma (in information_space) KL_ge_0: | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 373 | assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 374 | assumes ac: "absolutely_continuous \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 375 | assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 376 | shows "0 \<le> KL_divergence b M \<nu>" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 377 | using KL_eq_0 KL_gt_0[OF ps ac int] | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 378 | by (cases "\<forall>A\<in>sets M. \<nu> A = measure M A") (auto simp: le_less) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 379 | |
| 38656 | 380 | |
| 41833 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 381 | lemma (in sigma_finite_measure) KL_divergence_vimage: | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 382 | assumes T: "T \<in> measure_preserving M M'" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 383 | and T': "T' \<in> measure_preserving (M'\<lparr> measure := \<nu>' \<rparr>) (M\<lparr> measure := \<nu> \<rparr>)" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 384 | and inv: "\<And>x. x \<in> space M \<Longrightarrow> T' (T x) = x" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 385 | and inv': "\<And>x. x \<in> space M' \<Longrightarrow> T (T' x) = x" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 386 | and \<nu>': "measure_space (M'\<lparr>measure := \<nu>'\<rparr>)" "measure_space.absolutely_continuous M' \<nu>'" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 387 | and "1 < b" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 388 | shows "KL_divergence b M' \<nu>' = KL_divergence b M \<nu>" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 389 | proof - | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 390 | interpret \<nu>': measure_space "M'\<lparr>measure := \<nu>'\<rparr>" by fact | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 391 | have M: "measure_space (M\<lparr> measure := \<nu>\<rparr>)" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 392 | by (rule \<nu>'.measure_space_vimage[OF _ T'], simp) default | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 393 | have "sigma_algebra (M'\<lparr> measure := \<nu>'\<rparr>)" by default | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 394 | then have saM': "sigma_algebra M'" by simp | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 395 | then interpret M': measure_space M' by (rule measure_space_vimage) fact | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 396 | have ac: "absolutely_continuous \<nu>" unfolding absolutely_continuous_def | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 397 | proof safe | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 398 | fix N assume N: "N \<in> sets M" and N_0: "\<mu> N = 0" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 399 | then have N': "T' -` N \<inter> space M' \<in> sets M'" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 400 | using T' by (auto simp: measurable_def measure_preserving_def) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 401 | have "T -` (T' -` N \<inter> space M') \<inter> space M = N" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 402 | using inv T N sets_into_space[OF N] by (auto simp: measurable_def measure_preserving_def) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 403 | then have "measure M' (T' -` N \<inter> space M') = 0" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 404 | using measure_preservingD[OF T N'] N_0 by auto | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 405 | with \<nu>'(2) N' show "\<nu> N = 0" using measure_preservingD[OF T', of N] N | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 406 | unfolding M'.absolutely_continuous_def measurable_def by auto | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 407 | qed | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 408 | |
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 409 | have sa: "sigma_algebra (M\<lparr>measure := \<nu>\<rparr>)" by simp default | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 410 | have AE: "AE x. RN_deriv M' \<nu>' (T x) = RN_deriv M \<nu> x" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 411 | by (rule RN_deriv_vimage[OF T T' inv \<nu>']) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 412 | show ?thesis | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 413 | unfolding KL_divergence_def entropy_density_def comp_def | 
| 41833 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 414 | proof (subst \<nu>'.integral_vimage[OF sa T']) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 415 | show "(\<lambda>x. log b (real (RN_deriv M \<nu> x))) \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 416 | by (auto intro!: RN_deriv[OF M ac] borel_measurable_log[OF _ `1 < b`]) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 417 | have "(\<integral> x. log b (real (RN_deriv M' \<nu>' x)) \<partial>M'\<lparr>measure := \<nu>'\<rparr>) = | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 418 | (\<integral> x. log b (real (RN_deriv M' \<nu>' (T (T' x)))) \<partial>M'\<lparr>measure := \<nu>'\<rparr>)" (is "?l = _") | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 419 | using inv' by (auto intro!: \<nu>'.integral_cong) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 420 | also have "\<dots> = (\<integral> x. log b (real (RN_deriv M \<nu> (T' x))) \<partial>M'\<lparr>measure := \<nu>'\<rparr>)" (is "_ = ?r") | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 421 | using M ac AE | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 422 | by (intro \<nu>'.integral_cong_AE \<nu>'.almost_everywhere_vimage[OF sa T'] absolutely_continuous_AE[OF M]) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 423 | (auto elim!: AE_mp) | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 424 | finally show "?l = ?r" . | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 425 | qed | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 426 | qed | 
| 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 427 | |
| 40859 | 428 | lemma (in sigma_finite_measure) KL_divergence_cong: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 429 | assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?\<nu>") | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 430 | assumes [simp]: "sets N = sets M" "space N = space M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 431 | "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 432 | "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<nu>' A" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 433 | shows "KL_divergence b M \<nu> = KL_divergence b N \<nu>'" | 
| 40859 | 434 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 435 | interpret \<nu>: measure_space ?\<nu> by fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 436 | have "KL_divergence b M \<nu> = \<integral>x. log b (real (RN_deriv N \<nu>' x)) \<partial>?\<nu>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 437 | by (simp cong: RN_deriv_cong \<nu>.integral_cong add: KL_divergence_def entropy_density_def) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 438 | also have "\<dots> = KL_divergence b N \<nu>'" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 439 | by (auto intro!: \<nu>.integral_cong_measure[symmetric] simp: KL_divergence_def entropy_density_def comp_def) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 440 | finally show ?thesis . | 
| 40859 | 441 | qed | 
| 442 | ||
| 38656 | 443 | lemma (in finite_measure_space) KL_divergence_eq_finite: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 444 | assumes v: "finite_measure_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 40859 | 445 | assumes ac: "absolutely_continuous \<nu>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 446 |   shows "KL_divergence b M \<nu> = (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {x})))" (is "_ = ?sum")
 | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 447 | proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v] entropy_density_def) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 448 | interpret v: finite_measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 449 | have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 450 |   show "(\<Sum>x \<in> space M. log b (real (RN_deriv M \<nu> x)) * real (\<nu> {x})) = ?sum"
 | 
| 38656 | 451 | using RN_deriv_finite_measure[OF ms ac] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 452 | by (auto intro!: setsum_cong simp: field_simps) | 
| 38656 | 453 | qed | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 454 | |
| 38656 | 455 | lemma (in finite_prob_space) KL_divergence_positive_finite: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 456 | assumes v: "finite_prob_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 40859 | 457 | assumes ac: "absolutely_continuous \<nu>" | 
| 38656 | 458 | and "1 < b" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 459 | shows "0 \<le> KL_divergence b M \<nu>" | 
| 38656 | 460 | proof - | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 461 | interpret information_space M by default fact | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 462 | interpret v: finite_prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 463 | have ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)" by default | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 464 | from KL_ge_0[OF this ac v.integral_finite_singleton(1)] show ?thesis . | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 465 | qed | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 466 | |
| 39097 | 467 | subsection {* Mutual Information *}
 | 
| 468 | ||
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 469 | definition (in prob_space) | 
| 38656 | 470 | "mutual_information b S T X Y = | 
| 43920 | 471 | KL_divergence b (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>) | 
| 472 | (ereal\<circ>joint_distribution X Y)" | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 473 | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 474 | lemma (in information_space) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 475 | fixes S T X Y | 
| 43920 | 476 | defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 477 | shows "indep_var S X T Y \<longleftrightarrow> | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 478 | (random_variable S X \<and> random_variable T Y \<and> | 
| 43920 | 479 | measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y) \<and> | 
| 480 | integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>) | |
| 481 | (entropy_density b P (ereal\<circ>joint_distribution X Y)) \<and> | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 482 | mutual_information b S T X Y = 0)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 483 | proof safe | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 484 | assume indep: "indep_var S X T Y" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 485 | then have "random_variable S X" "random_variable T Y" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 486 | by (blast dest: indep_var_rv1 indep_var_rv2)+ | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 487 | then show "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 488 | by blast+ | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 489 | |
| 43920 | 490 | interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 491 | by (rule distribution_prob_space) fact | 
| 43920 | 492 | interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 493 | by (rule distribution_prob_space) fact | 
| 43920 | 494 | interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 495 | interpret XY: information_space XY.P b by default (rule b_gt_1) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 496 | |
| 43920 | 497 | let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 498 |   { fix A assume "A \<in> sets XY.P"
 | 
| 43920 | 499 | then have "ereal (joint_distribution X Y A) = XY.\<mu> A" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 500 | using indep_var_distributionD[OF indep] | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 501 | by (simp add: XY.P.finite_measure_eq) } | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 502 | note j_eq = this | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 503 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 504 | interpret J: prob_space ?J | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 505 | using j_eq by (intro XY.prob_space_cong) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 506 | |
| 43920 | 507 | have ac: "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 508 | by (simp add: XY.absolutely_continuous_def j_eq) | 
| 43920 | 509 | then show "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 510 | unfolding P_def . | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 511 | |
| 43920 | 512 | have ed: "entropy_density b XY.P (ereal\<circ>joint_distribution X Y) \<in> borel_measurable XY.P" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 513 | by (rule XY.measurable_entropy_density) (default | fact)+ | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 514 | |
| 43920 | 515 | have "AE x in XY.P. 1 = RN_deriv XY.P (ereal\<circ>joint_distribution X Y) x" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 516 | proof (rule XY.RN_deriv_unique[OF _ ac]) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 517 | show "measure_space ?J" by default | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 518 | fix A assume "A \<in> sets XY.P" | 
| 43920 | 519 | then show "(ereal\<circ>joint_distribution X Y) A = (\<integral>\<^isup>+ x. 1 * indicator A x \<partial>XY.P)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 520 | by (simp add: j_eq) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 521 | qed (insert XY.measurable_const[of 1 borel], auto) | 
| 43920 | 522 | then have ae_XY: "AE x in XY.P. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 523 | by (elim XY.AE_mp) (simp add: entropy_density_def) | 
| 43920 | 524 | have ae_J: "AE x in ?J. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 525 | proof (rule XY.absolutely_continuous_AE) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 526 | show "measure_space ?J" by default | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 527 | show "XY.absolutely_continuous (measure ?J)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 528 | using ac by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 529 | qed (insert ae_XY, simp_all) | 
| 43920 | 530 | then show "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>) | 
| 531 | (entropy_density b P (ereal\<circ>joint_distribution X Y))" | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 532 | unfolding P_def | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 533 | using ed XY.measurable_const[of 0 borel] | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 534 | by (subst J.integrable_cong_AE) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 535 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 536 | show "mutual_information b S T X Y = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 537 | unfolding mutual_information_def KL_divergence_def P_def | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 538 | by (subst J.integral_cong_AE[OF ae_J]) simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 539 | next | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 540 | assume "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 541 | then have rvs: "random_variable S X" "random_variable T Y" by blast+ | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 542 | |
| 43920 | 543 | interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 544 | by (rule distribution_prob_space) fact | 
| 43920 | 545 | interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 546 | by (rule distribution_prob_space) fact | 
| 43920 | 547 | interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 548 | interpret XY: information_space XY.P b by default (rule b_gt_1) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 549 | |
| 43920 | 550 | let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 551 | interpret J: prob_space ?J | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 552 | using rvs by (intro joint_distribution_prob_space) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 553 | |
| 43920 | 554 | assume ac: "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)" | 
| 555 | assume int: "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>) | |
| 556 | (entropy_density b P (ereal\<circ>joint_distribution X Y))" | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 557 | assume I_eq_0: "mutual_information b S T X Y = 0" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 558 | |
| 43920 | 559 | have eq: "\<forall>A\<in>sets XY.P. (ereal \<circ> joint_distribution X Y) A = XY.\<mu> A" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 560 | proof (rule XY.KL_eq_0_imp) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 561 | show "prob_space ?J" by default | 
| 43920 | 562 | show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 563 | using ac by (simp add: P_def) | 
| 43920 | 564 | show "integrable ?J (entropy_density b XY.P (ereal\<circ>joint_distribution X Y))" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 565 | using int by (simp add: P_def) | 
| 43920 | 566 | show "KL_divergence b XY.P (ereal\<circ>joint_distribution X Y) = 0" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 567 | using I_eq_0 unfolding mutual_information_def by (simp add: P_def) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 568 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 569 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 570 |   { fix S X assume "sigma_algebra S"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 571 | interpret S: sigma_algebra S by fact | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 572 |     have "Int_stable \<lparr>space = space M, sets = {X -` A \<inter> space M |A. A \<in> sets S}\<rparr>"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 573 | proof (safe intro!: Int_stableI) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 574 | fix A B assume "A \<in> sets S" "B \<in> sets S" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 575 | 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" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 576 | by (intro exI[of _ "A \<inter> B"]) auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 577 | qed } | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 578 | note Int_stable = this | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 579 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 580 | show "indep_var S X T Y" unfolding indep_var_eq | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 581 | proof (intro conjI indep_set_sigma_sets Int_stable) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 582 |     show "indep_set {X -` A \<inter> space M |A. A \<in> sets S} {Y -` A \<inter> space M |A. A \<in> sets T}"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 583 | proof (safe intro!: indep_setI) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 584 |       { fix A assume "A \<in> sets S" then show "X -` A \<inter> space M \<in> sets M"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 585 | using `X \<in> measurable M S` by (auto intro: measurable_sets) } | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 586 |       { fix A assume "A \<in> sets T" then show "Y -` A \<inter> space M \<in> sets M"
 | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 587 | using `Y \<in> measurable M T` by (auto intro: measurable_sets) } | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 588 | next | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 589 | fix A B assume ab: "A \<in> sets S" "B \<in> sets T" | 
| 43920 | 590 | have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) = | 
| 591 | ereal (joint_distribution X Y (A \<times> B))" | |
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 592 | unfolding distribution_def | 
| 43920 | 593 | by (intro arg_cong[where f="\<lambda>C. ereal (prob C)"]) auto | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 594 | also have "\<dots> = XY.\<mu> (A \<times> B)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 595 | using ab eq by (auto simp: XY.finite_measure_eq) | 
| 43920 | 596 | also have "\<dots> = ereal (distribution X A) * ereal (distribution Y B)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 597 | using ab by (simp add: XY.pair_measure_times) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 598 | finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) = | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 599 | prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 600 | unfolding distribution_def by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 601 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 602 | qed fact+ | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 603 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 604 | |
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 605 | lemma (in information_space) mutual_information_commute_generic: | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 606 | assumes X: "random_variable S X" and Y: "random_variable T Y" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 607 | assumes ac: "measure_space.absolutely_continuous | 
| 43920 | 608 | (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>) (ereal\<circ>joint_distribution X Y)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 609 | shows "mutual_information b S T X Y = mutual_information b T S Y X" | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 610 | proof - | 
| 43920 | 611 | let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 612 | interpret S: prob_space ?S using X by (rule distribution_prob_space) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 613 | interpret T: prob_space ?T using Y by (rule distribution_prob_space) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 614 | interpret P: pair_prob_space ?S ?T .. | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 615 | interpret Q: pair_prob_space ?T ?S .. | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 616 | show ?thesis | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 617 | unfolding mutual_information_def | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 618 | proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1]) | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 619 | show "(\<lambda>(x,y). (y,x)) \<in> measure_preserving | 
| 43920 | 620 | (P.P \<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := ereal\<circ>joint_distribution Y X\<rparr>)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 621 | using X Y unfolding measurable_def | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 622 | unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 623 | by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>']) | 
| 43920 | 624 | have "prob_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 625 | using X Y by (auto intro!: distribution_prob_space random_variable_pairI) | 
| 43920 | 626 | then show "measure_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)" | 
| 43340 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 627 | unfolding prob_space_def by simp | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 628 | qed auto | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 629 | qed | 
| 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 hoelzl parents: 
42148diff
changeset | 630 | |
| 40859 | 631 | definition (in prob_space) | 
| 632 | "entropy b s X = mutual_information b s s X X" | |
| 633 | ||
| 634 | abbreviation (in information_space) | |
| 635 |   mutual_information_Pow ("\<I>'(_ ; _')") where
 | |
| 36624 | 636 | "\<I>(X ; Y) \<equiv> mutual_information b | 
| 43920 | 637 | \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr> | 
| 638 | \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 639 | |
| 40859 | 640 | lemma (in prob_space) finite_variables_absolutely_continuous: | 
| 641 | assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 642 | shows "measure_space.absolutely_continuous | 
| 43920 | 643 | (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>) | 
| 644 | (ereal\<circ>joint_distribution X Y)" | |
| 40859 | 645 | proof - | 
| 43920 | 646 | interpret X: finite_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 647 | using X by (rule distribution_finite_prob_space) | 
| 43920 | 648 | interpret Y: finite_prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 649 | using Y by (rule distribution_finite_prob_space) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 650 | interpret XY: pair_finite_prob_space | 
| 43920 | 651 | "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr> measure := ereal\<circ>distribution Y\<rparr>" by default | 
| 652 | interpret P: finite_prob_space "XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 653 | using assms by (auto intro!: joint_distribution_finite_prob_space) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 654 | note rv = assms[THEN finite_random_variableD] | 
| 43920 | 655 | show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)" | 
| 40859 | 656 | proof (rule XY.absolutely_continuousI) | 
| 43920 | 657 | show "finite_measure_space (XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 658 |     fix x assume "x \<in> space XY.P" and "XY.\<mu> {x} = 0"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 659 | then obtain a b where "x = (a, b)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 660 |       and "distribution X {a} = 0 \<or> distribution Y {b} = 0"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 661 | by (cases x) (auto simp: space_pair_measure) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 662 | with finite_distribution_order(5,6)[OF X Y] | 
| 43920 | 663 |     show "(ereal \<circ> joint_distribution X Y) {x} = 0" by auto
 | 
| 40859 | 664 | qed | 
| 665 | qed | |
| 666 | ||
| 667 | lemma (in information_space) | |
| 668 | assumes MX: "finite_random_variable MX X" | |
| 669 | assumes MY: "finite_random_variable MY Y" | |
| 670 | shows mutual_information_generic_eq: | |
| 36624 | 671 | "mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY. | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 672 |       joint_distribution X Y {(x,y)} *
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 673 |       log b (joint_distribution X Y {(x,y)} /
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 674 |       (distribution X {x} * distribution Y {y})))"
 | 
| 40859 | 675 | (is ?sum) | 
| 36624 | 676 | and mutual_information_positive_generic: | 
| 40859 | 677 | "0 \<le> mutual_information b MX MY X Y" (is ?positive) | 
| 36624 | 678 | proof - | 
| 43920 | 679 | interpret X: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 680 | using MX by (rule distribution_finite_prob_space) | 
| 43920 | 681 | interpret Y: finite_prob_space "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 682 | using MY by (rule distribution_finite_prob_space) | 
| 43920 | 683 | interpret XY: pair_finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>" "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default | 
| 684 | interpret P: finite_prob_space "XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 685 | using assms by (auto intro!: joint_distribution_finite_prob_space) | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 686 | |
| 43920 | 687 | have P_ms: "finite_measure_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default | 
| 688 | have P_ps: "finite_prob_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default | |
| 36624 | 689 | |
| 40859 | 690 | show ?sum | 
| 38656 | 691 | unfolding Let_def mutual_information_def | 
| 40859 | 692 | by (subst XY.KL_divergence_eq_finite[OF P_ms finite_variables_absolutely_continuous[OF MX MY]]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 693 | (auto simp add: space_pair_measure setsum_cartesian_product') | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 694 | |
| 36624 | 695 | show ?positive | 
| 40859 | 696 | using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1] | 
| 697 | unfolding mutual_information_def . | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 698 | qed | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 699 | |
| 41661 | 700 | lemma (in information_space) mutual_information_commute: | 
| 701 | assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y" | |
| 702 | shows "mutual_information b S T X Y = mutual_information b T S Y X" | |
| 703 | unfolding mutual_information_generic_eq[OF X Y] mutual_information_generic_eq[OF Y X] | |
| 704 | unfolding joint_distribution_commute_singleton[of X Y] | |
| 705 | by (auto simp add: ac_simps intro!: setsum_reindex_cong[OF swap_inj_on]) | |
| 706 | ||
| 707 | lemma (in information_space) mutual_information_commute_simple: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 708 | assumes X: "simple_function M X" and Y: "simple_function M Y" | 
| 41661 | 709 | shows "\<I>(X;Y) = \<I>(Y;X)" | 
| 41833 
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
 hoelzl parents: 
41689diff
changeset | 710 | by (intro mutual_information_commute X Y simple_function_imp_finite_random_variable) | 
| 41661 | 711 | |
| 40859 | 712 | lemma (in information_space) mutual_information_eq: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 713 | assumes "simple_function M X" "simple_function M Y" | 
| 40859 | 714 | shows "\<I>(X;Y) = (\<Sum> (x,y) \<in> X ` space M \<times> Y ` space M. | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 715 |     distribution (\<lambda>x. (X x, Y x)) {(x,y)} * log b (distribution (\<lambda>x. (X x, Y x)) {(x,y)} /
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 716 |                                                    (distribution X {x} * distribution Y {y})))"
 | 
| 40859 | 717 | using assms by (simp add: mutual_information_generic_eq) | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 718 | |
| 40859 | 719 | lemma (in information_space) mutual_information_generic_cong: | 
| 39097 | 720 | assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" | 
| 721 | assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" | |
| 40859 | 722 | shows "mutual_information b MX MY X Y = mutual_information b MX MY X' Y'" | 
| 723 | unfolding mutual_information_def using X Y | |
| 724 | by (simp cong: distribution_cong) | |
| 39097 | 725 | |
| 40859 | 726 | lemma (in information_space) mutual_information_cong: | 
| 727 | assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" | |
| 728 | assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" | |
| 729 | shows "\<I>(X; Y) = \<I>(X'; Y')" | |
| 730 | unfolding mutual_information_def using X Y | |
| 731 | by (simp cong: distribution_cong image_cong) | |
| 732 | ||
| 733 | lemma (in information_space) mutual_information_positive: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 734 | assumes "simple_function M X" "simple_function M Y" | 
| 40859 | 735 | shows "0 \<le> \<I>(X;Y)" | 
| 736 | using assms by (simp add: mutual_information_positive_generic) | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 737 | |
| 39097 | 738 | subsection {* Entropy *}
 | 
| 739 | ||
| 40859 | 740 | abbreviation (in information_space) | 
| 741 |   entropy_Pow ("\<H>'(_')") where
 | |
| 43920 | 742 | "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr> X" | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 743 | |
| 40859 | 744 | lemma (in information_space) entropy_generic_eq: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 745 | fixes X :: "'a \<Rightarrow> 'c" | 
| 40859 | 746 | assumes MX: "finite_random_variable MX X" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 747 |   shows "entropy b MX X = -(\<Sum> x \<in> space MX. distribution X {x} * log b (distribution X {x}))"
 | 
| 39097 | 748 | proof - | 
| 43920 | 749 | interpret MX: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 750 | using MX by (rule distribution_finite_prob_space) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 751 |   let "?X x" = "distribution X {x}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 752 |   let "?XX x y" = "joint_distribution X X {(x, y)}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 753 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 754 |   { fix x y :: 'c
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 755 |     { assume "x \<noteq> y"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 756 |       then have "(\<lambda>x. (X x, X x)) -` {(x,y)} \<inter> space M = {}" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 757 |       then have "joint_distribution X X {(x, y)} = 0" by (simp add: distribution_def) }
 | 
| 39097 | 758 | then have "?XX x y * log b (?XX x y / (?X x * ?X y)) = | 
| 759 | (if x = y then - ?X y * log b (?X y) else 0)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 760 | by (auto simp: log_simps zero_less_mult_iff) } | 
| 39097 | 761 | note remove_XX = this | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 762 | |
| 39097 | 763 | show ?thesis | 
| 764 | unfolding entropy_def mutual_information_generic_eq[OF MX MX] | |
| 765 | unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 766 | using MX.finite_space by (auto simp: setsum_cases) | 
| 39097 | 767 | qed | 
| 36624 | 768 | |
| 40859 | 769 | lemma (in information_space) entropy_eq: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 770 | assumes "simple_function M X" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 771 |   shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. distribution X {x} * log b (distribution X {x}))"
 | 
| 40859 | 772 | using assms by (simp add: entropy_generic_eq) | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 773 | |
| 40859 | 774 | lemma (in information_space) entropy_positive: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 775 | "simple_function M X \<Longrightarrow> 0 \<le> \<H>(X)" | 
| 40859 | 776 | unfolding entropy_def by (simp add: mutual_information_positive) | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 777 | |
| 40859 | 778 | lemma (in information_space) entropy_certainty_eq_0: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 779 |   assumes X: "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
 | 
| 39097 | 780 | shows "\<H>(X) = 0" | 
| 781 | proof - | |
| 43920 | 782 | let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal\<circ>distribution X\<rparr>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 783 | note simple_function_imp_finite_random_variable[OF `simple_function M X`] | 
| 43920 | 784 | from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"] | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 785 | interpret X: finite_prob_space ?X by simp | 
| 39097 | 786 |   have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
 | 
| 787 |     using X.measure_compl[of "{x}"] assms by auto
 | |
| 788 | also have "\<dots> = 0" using X.prob_space assms by auto | |
| 789 |   finally have X0: "distribution X (X ` space M - {x}) = 0" by auto
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 790 |   { fix y assume *: "y \<in> X ` space M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 791 |     { assume asm: "y \<noteq> x"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 792 |       with * have "{y} \<subseteq> X ` space M - {x}" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 793 | from X.measure_mono[OF this] X0 asm * | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 794 |       have "distribution X {y} = 0"  by (auto intro: antisym) }
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 795 |     then have "distribution X {y} = (if x = y then 1 else 0)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 796 | using assms by auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 797 | note fi = this | 
| 39097 | 798 | have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 799 | show ?thesis unfolding entropy_eq[OF `simple_function M X`] by (auto simp: y fi) | 
| 39097 | 800 | qed | 
| 801 | ||
| 40859 | 802 | lemma (in information_space) entropy_le_card_not_0: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 803 | assumes X: "simple_function M X" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 804 |   shows "\<H>(X) \<le> log b (card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}))"
 | 
| 39097 | 805 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 806 |   let "?p x" = "distribution X {x}"
 | 
| 39097 | 807 | have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 808 | unfolding entropy_eq[OF X] setsum_negf[symmetric] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 809 | by (auto intro!: setsum_cong simp: log_simps) | 
| 39097 | 810 | also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 811 | using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution[OF X] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 812 | by (intro log_setsum') (auto simp: simple_function_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 813 | also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?p x \<noteq> 0 then 1 else 0)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 814 | by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) auto | 
| 39097 | 815 | finally show ?thesis | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 816 | using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def) | 
| 39097 | 817 | qed | 
| 818 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 819 | lemma (in prob_space) measure'_translate: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 820 | assumes X: "random_variable S X" and A: "A \<in> sets S" | 
| 43920 | 821 | shows "finite_measure.\<mu>' (S\<lparr> measure := ereal\<circ>distribution X \<rparr>) A = distribution X A" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 822 | proof - | 
| 43920 | 823 | interpret S: prob_space "S\<lparr> measure := ereal\<circ>distribution X \<rparr>" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 824 | using distribution_prob_space[OF X] . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 825 | from A show "S.\<mu>' A = distribution X A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 826 | unfolding S.\<mu>'_def by (simp add: distribution_def_raw \<mu>'_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 827 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 828 | |
| 40859 | 829 | lemma (in information_space) entropy_uniform_max: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 830 | assumes X: "simple_function M X" | 
| 39097 | 831 |   assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
 | 
| 832 | shows "\<H>(X) = log b (real (card (X ` space M)))" | |
| 833 | proof - | |
| 43920 | 834 | let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := ereal\<circ>distribution X\<rparr>" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 835 | note frv = simple_function_imp_finite_random_variable[OF X] | 
| 43920 | 836 | from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"] | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 837 | interpret X: finite_prob_space ?X by simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 838 | note rv = finite_random_variableD[OF frv] | 
| 39097 | 839 | have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 840 | using `simple_function M X` not_empty by (auto simp: simple_function_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 841 |   { fix x assume "x \<in> space ?X"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 842 |     moreover then have "X.\<mu>' {x} = 1 / card (space ?X)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 843 | proof (rule X.uniform_prob) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 844 | fix x y assume "x \<in> space ?X" "y \<in> space ?X" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 845 |       with assms(2)[of x y] show "X.\<mu>' {x} = X.\<mu>' {y}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 846 | by (subst (1 2) measure'_translate[OF rv]) auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 847 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 848 |     ultimately have "distribution X {x} = 1 / card (space ?X)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 849 | by (subst (asm) measure'_translate[OF rv]) auto } | 
| 39097 | 850 | thus ?thesis | 
| 40859 | 851 | using not_empty X.finite_space b_gt_1 card_gt0 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 852 | by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps) | 
| 39097 | 853 | qed | 
| 854 | ||
| 40859 | 855 | lemma (in information_space) entropy_le_card: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 856 | assumes "simple_function M X" | 
| 40859 | 857 | shows "\<H>(X) \<le> log b (real (card (X ` space M)))" | 
| 39097 | 858 | proof cases | 
| 859 |   assume "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} = {}"
 | |
| 860 |   then have "\<And>x. x\<in>X`space M \<Longrightarrow> distribution X {x} = 0" by auto
 | |
| 861 | moreover | |
| 862 | have "0 < card (X`space M)" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 863 | using `simple_function M X` not_empty | 
| 40859 | 864 | by (auto simp: card_gt_0_iff simple_function_def) | 
| 39097 | 865 | then have "log b 1 \<le> log b (real (card (X`space M)))" | 
| 866 | using b_gt_1 by (intro log_le) auto | |
| 40859 | 867 | ultimately show ?thesis using assms by (simp add: entropy_eq) | 
| 39097 | 868 | next | 
| 869 |   assume False: "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} \<noteq> {}"
 | |
| 870 |   have "card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}) \<le> card (X ` space M)"
 | |
| 40859 | 871 | (is "?A \<le> ?B") using assms not_empty by (auto intro!: card_mono simp: simple_function_def) | 
| 872 | note entropy_le_card_not_0[OF assms] | |
| 39097 | 873 | also have "log b (real ?A) \<le> log b (real ?B)" | 
| 40859 | 874 | using b_gt_1 False not_empty `?A \<le> ?B` assms | 
| 875 | by (auto intro!: log_le simp: card_gt_0_iff simp: simple_function_def) | |
| 39097 | 876 | finally show ?thesis . | 
| 877 | qed | |
| 878 | ||
| 40859 | 879 | lemma (in information_space) entropy_commute: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 880 | assumes "simple_function M X" "simple_function M Y" | 
| 40859 | 881 | shows "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))" | 
| 39097 | 882 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 883 | have sf: "simple_function M (\<lambda>x. (X x, Y x))" "simple_function M (\<lambda>x. (Y x, X x))" | 
| 40859 | 884 | using assms by (auto intro: simple_function_Pair) | 
| 39097 | 885 | have *: "(\<lambda>x. (Y x, X x))`space M = (\<lambda>(a,b). (b,a))`(\<lambda>x. (X x, Y x))`space M" | 
| 886 | by auto | |
| 887 | have inj: "\<And>X. inj_on (\<lambda>(a,b). (b,a)) X" | |
| 888 | by (auto intro!: inj_onI) | |
| 889 | show ?thesis | |
| 40859 | 890 | unfolding sf[THEN entropy_eq] unfolding * setsum_reindex[OF inj] | 
| 39097 | 891 | by (simp add: joint_distribution_commute[of Y X] split_beta) | 
| 892 | qed | |
| 893 | ||
| 40859 | 894 | lemma (in information_space) entropy_eq_cartesian_product: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 895 | assumes "simple_function M X" "simple_function M Y" | 
| 40859 | 896 | shows "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M. | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 897 |     joint_distribution X Y {(x,y)} * log b (joint_distribution X Y {(x,y)}))"
 | 
| 39097 | 898 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 899 | have sf: "simple_function M (\<lambda>x. (X x, Y x))" | 
| 40859 | 900 | using assms by (auto intro: simple_function_Pair) | 
| 39097 | 901 |   { fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M"
 | 
| 902 |     then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto
 | |
| 903 |     then have "joint_distribution X Y {x} = 0"
 | |
| 904 | unfolding distribution_def by auto } | |
| 40859 | 905 | then show ?thesis using sf assms | 
| 906 | unfolding entropy_eq[OF sf] neg_equal_iff_equal setsum_cartesian_product | |
| 907 | by (auto intro!: setsum_mono_zero_cong_left simp: simple_function_def) | |
| 39097 | 908 | qed | 
| 909 | ||
| 910 | subsection {* Conditional Mutual Information *}
 | |
| 911 | ||
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 912 | definition (in prob_space) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 913 | "conditional_mutual_information b MX MY MZ X Y Z \<equiv> | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 914 | mutual_information b MX (MY \<Otimes>\<^isub>M MZ) X (\<lambda>x. (Y x, Z x)) - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 915 | mutual_information b MX MZ X Z" | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 916 | |
| 40859 | 917 | abbreviation (in information_space) | 
| 918 |   conditional_mutual_information_Pow ("\<I>'( _ ; _ | _ ')") where
 | |
| 36624 | 919 | "\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b | 
| 43920 | 920 | \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr> | 
| 921 | \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> | |
| 922 | \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = ereal\<circ>distribution Z \<rparr> | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 923 | X Y Z" | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 924 | |
| 40859 | 925 | lemma (in information_space) conditional_mutual_information_generic_eq: | 
| 926 | assumes MX: "finite_random_variable MX X" | |
| 927 | and MY: "finite_random_variable MY Y" | |
| 928 | and MZ: "finite_random_variable MZ Z" | |
| 929 | shows "conditional_mutual_information b MX MY MZ X Y Z = (\<Sum>(x, y, z) \<in> space MX \<times> space MY \<times> space MZ. | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 930 |              distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 931 |              log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 932 |     (joint_distribution X Z {(x, z)} * (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 933 | (is "_ = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z))))") | 
| 40859 | 934 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 935 |   let ?X = "\<lambda>x. distribution X {x}"
 | 
| 40859 | 936 | note finite_var = MX MY MZ | 
| 937 | note YZ = finite_random_variable_pairI[OF finite_var(2,3)] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 938 | note XYZ = finite_random_variable_pairI[OF MX YZ] | 
| 40859 | 939 | note XZ = finite_random_variable_pairI[OF finite_var(1,3)] | 
| 940 | note ZX = finite_random_variable_pairI[OF finite_var(3,1)] | |
| 941 | note YZX = finite_random_variable_pairI[OF finite_var(2) ZX] | |
| 942 | note order1 = | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 943 | finite_distribution_order(5,6)[OF finite_var(1) YZ] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 944 | finite_distribution_order(5,6)[OF finite_var(1,3)] | 
| 40859 | 945 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 946 | note random_var = finite_var[THEN finite_random_variableD] | 
| 40859 | 947 | note finite = finite_var(1) YZ finite_var(3) XZ YZX | 
| 948 | ||
| 949 |   have order2: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
 | |
| 950 |           \<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
 | |
| 951 | unfolding joint_distribution_commute_singleton[of X] | |
| 952 | unfolding joint_distribution_assoc_singleton[symmetric] | |
| 953 | using finite_distribution_order(6)[OF finite_var(2) ZX] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 954 | by auto | 
| 36624 | 955 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 956 | have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z)))) = | 
| 40859 | 957 | (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * (log b (?XYZ x y z / (?X x * ?YZ y z)) - log b (?XZ x z / (?X x * ?Z z))))" | 
| 958 | (is "(\<Sum>(x, y, z)\<in>?S. ?L x y z) = (\<Sum>(x, y, z)\<in>?S. ?R x y z)") | |
| 959 | proof (safe intro!: setsum_cong) | |
| 960 | fix x y z assume space: "x \<in> space MX" "y \<in> space MY" "z \<in> space MZ" | |
| 961 | show "?L x y z = ?R x y z" | |
| 962 | proof cases | |
| 963 | assume "?XYZ x y z \<noteq> 0" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 964 | with space have "0 < ?X x" "0 < ?Z z" "0 < ?XZ x z" "0 < ?YZ y z" "0 < ?XYZ x y z" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 965 | using order1 order2 by (auto simp: less_le) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 966 | with b_gt_1 show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 967 | by (simp add: log_mult log_divide zero_less_mult_iff zero_less_divide_iff) | 
| 40859 | 968 | qed simp | 
| 969 | qed | |
| 970 | also have "\<dots> = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) - | |
| 971 | (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XZ x z / (?X x * ?Z z)))" | |
| 972 | by (auto simp add: setsum_subtractf[symmetric] field_simps intro!: setsum_cong) | |
| 973 | also have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XZ x z / (?X x * ?Z z))) = | |
| 974 | (\<Sum>(x, z)\<in>space MX \<times> space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z)))" | |
| 975 | unfolding setsum_cartesian_product[symmetric] setsum_commute[of _ _ "space MY"] | |
| 976 | setsum_left_distrib[symmetric] | |
| 977 | unfolding joint_distribution_commute_singleton[of X] | |
| 978 | unfolding joint_distribution_assoc_singleton[symmetric] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 979 | using setsum_joint_distribution_singleton[OF finite_var(2) ZX] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 980 | by (intro setsum_cong refl) (simp add: space_pair_measure) | 
| 40859 | 981 | also have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) - | 
| 982 | (\<Sum>(x, z)\<in>space MX \<times> space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z))) = | |
| 983 | conditional_mutual_information b MX MY MZ X Y Z" | |
| 984 | unfolding conditional_mutual_information_def | |
| 985 | unfolding mutual_information_generic_eq[OF finite_var(1,3)] | |
| 986 | unfolding mutual_information_generic_eq[OF finite_var(1) YZ] | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 987 | by (simp add: space_sigma space_pair_measure setsum_cartesian_product') | 
| 40859 | 988 | finally show ?thesis by simp | 
| 989 | qed | |
| 990 | ||
| 991 | lemma (in information_space) conditional_mutual_information_eq: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 992 | assumes "simple_function M X" "simple_function M Y" "simple_function M Z" | 
| 40859 | 993 | shows "\<I>(X;Y|Z) = (\<Sum>(x, y, z) \<in> X`space M \<times> Y`space M \<times> Z`space M. | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 994 |              distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 995 |              log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 996 |     (joint_distribution X Z {(x, z)} * joint_distribution Y Z {(y,z)} / distribution Z {z})))"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 997 | by (subst conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 998 | simp | 
| 40859 | 999 | |
| 1000 | lemma (in information_space) conditional_mutual_information_eq_mutual_information: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1001 | assumes X: "simple_function M X" and Y: "simple_function M Y" | 
| 40859 | 1002 | shows "\<I>(X ; Y) = \<I>(X ; Y | (\<lambda>x. ()))" | 
| 36624 | 1003 | proof - | 
| 1004 |   have [simp]: "(\<lambda>x. ()) ` space M = {()}" using not_empty by auto
 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1005 | have C: "simple_function M (\<lambda>x. ())" by auto | 
| 36624 | 1006 | show ?thesis | 
| 40859 | 1007 | unfolding conditional_mutual_information_eq[OF X Y C] | 
| 1008 | unfolding mutual_information_eq[OF X Y] | |
| 36624 | 1009 | by (simp add: setsum_cartesian_product' distribution_remove_const) | 
| 1010 | qed | |
| 1011 | ||
| 40859 | 1012 | lemma (in prob_space) distribution_unit[simp]: "distribution (\<lambda>x. ()) {()} = 1"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1013 | unfolding distribution_def using prob_space by auto | 
| 40859 | 1014 | |
| 1015 | lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\<lambda>x. (X x, ())) {(a, ())} = distribution X {a}"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1016 | unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) | 
| 40859 | 1017 | |
| 1018 | lemma (in prob_space) setsum_distribution: | |
| 1019 |   assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
 | |
| 1020 |   using setsum_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV \<rparr>" "\<lambda>x. ()" "{()}"]
 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1021 | using sigma_algebra_Pow[of "UNIV::unit set" "()"] by simp | 
| 40859 | 1022 | |
| 1023 | lemma (in prob_space) setsum_real_distribution: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1024 |   fixes MX :: "('c, 'd) measure_space_scheme"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1025 |   assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1026 |   using setsum_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV, measure = undefined \<rparr>" "\<lambda>x. ()" "{()}"]
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1027 | using sigma_algebra_Pow[of "UNIV::unit set" "\<lparr> measure = undefined \<rparr>"] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1028 | by auto | 
| 40859 | 1029 | |
| 1030 | lemma (in information_space) conditional_mutual_information_generic_positive: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1031 | assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" and Z: "finite_random_variable MZ Z" | 
| 40859 | 1032 | shows "0 \<le> conditional_mutual_information b MX MY MZ X Y Z" | 
| 1033 | proof (cases "space MX \<times> space MY \<times> space MZ = {}")
 | |
| 1034 | case True show ?thesis | |
| 1035 | unfolding conditional_mutual_information_generic_eq[OF assms] True | |
| 1036 | by simp | |
| 1037 | next | |
| 1038 | case False | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1039 | let ?dXYZ = "distribution (\<lambda>x. (X x, Y x, Z x))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1040 | let ?dXZ = "joint_distribution X Z" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1041 | let ?dYZ = "joint_distribution Y Z" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1042 | let ?dX = "distribution X" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1043 | let ?dZ = "distribution Z" | 
| 40859 | 1044 | let ?M = "space MX \<times> space MY \<times> space MZ" | 
| 36624 | 1045 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1046 | note YZ = finite_random_variable_pairI[OF Y Z] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1047 | note XZ = finite_random_variable_pairI[OF X Z] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1048 | note ZX = finite_random_variable_pairI[OF Z X] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1049 | note YZ = finite_random_variable_pairI[OF Y Z] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1050 | note XYZ = finite_random_variable_pairI[OF X YZ] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1051 | note finite = Z YZ XZ XYZ | 
| 40859 | 1052 |   have order: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
 | 
| 1053 |           \<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
 | |
| 1054 | unfolding joint_distribution_commute_singleton[of X] | |
| 1055 | unfolding joint_distribution_assoc_singleton[symmetric] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1056 | using finite_distribution_order(6)[OF Y ZX] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1057 | by auto | 
| 40859 | 1058 | |
| 1059 | note order = order | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1060 | finite_distribution_order(5,6)[OF X YZ] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1061 | finite_distribution_order(5,6)[OF Y Z] | 
| 40859 | 1062 | |
| 1063 |   have "- conditional_mutual_information b MX MY MZ X Y Z = - (\<Sum>(x, y, z) \<in> ?M. ?dXYZ {(x, y, z)} *
 | |
| 1064 |     log b (?dXYZ {(x, y, z)} / (?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})))"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1065 | unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal by auto | 
| 40859 | 1066 |   also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1067 | unfolding split_beta' | 
| 36624 | 1068 | proof (rule log_setsum_divide) | 
| 40859 | 1069 |     show "?M \<noteq> {}" using False by simp
 | 
| 36624 | 1070 | show "1 < b" using b_gt_1 . | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1071 | |
| 40859 | 1072 | show "finite ?M" using assms | 
| 1073 | unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by auto | |
| 1074 | ||
| 1075 |     show "(\<Sum>x\<in>?M. ?dXYZ {(fst x, fst (snd x), snd (snd x))}) = 1"
 | |
| 1076 | unfolding setsum_cartesian_product' | |
| 1077 | unfolding setsum_commute[of _ "space MY"] | |
| 1078 | unfolding setsum_commute[of _ "space MZ"] | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1079 | by (simp_all add: space_pair_measure | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1080 | setsum_joint_distribution_singleton[OF X YZ] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1081 | setsum_joint_distribution_singleton[OF Y Z] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1082 | setsum_distribution[OF Z]) | 
| 40859 | 1083 | |
| 36624 | 1084 | fix x assume "x \<in> ?M" | 
| 38656 | 1085 | let ?x = "(fst x, fst (snd x), snd (snd x))" | 
| 1086 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1087 |     show "0 \<le> ?dXYZ {?x}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1088 |       "0 \<le> ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1089 | by (simp_all add: mult_nonneg_nonneg divide_nonneg_nonneg) | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1090 | |
| 38656 | 1091 |     assume *: "0 < ?dXYZ {?x}"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1092 |     with `x \<in> ?M` finite order show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1093 | by (cases x) (auto simp add: zero_le_mult_iff zero_le_divide_iff less_le) | 
| 40859 | 1094 | qed | 
| 1095 |   also have "(\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\<Sum>z\<in>space MZ. ?dZ {z})"
 | |
| 36624 | 1096 | apply (simp add: setsum_cartesian_product') | 
| 1097 | apply (subst setsum_commute) | |
| 1098 | apply (subst (2) setsum_commute) | |
| 40859 | 1099 | by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1100 | setsum_joint_distribution_singleton[OF X Z] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1101 | setsum_joint_distribution_singleton[OF Y Z] | 
| 36624 | 1102 | intro!: setsum_cong) | 
| 40859 | 1103 |   also have "log b (\<Sum>z\<in>space MZ. ?dZ {z}) = 0"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1104 | unfolding setsum_real_distribution[OF Z] by simp | 
| 40859 | 1105 | finally show ?thesis by simp | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1106 | qed | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1107 | |
| 40859 | 1108 | lemma (in information_space) conditional_mutual_information_positive: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1109 | assumes "simple_function M X" and "simple_function M Y" and "simple_function M Z" | 
| 40859 | 1110 | shows "0 \<le> \<I>(X;Y|Z)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1111 | by (rule conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]]) | 
| 40859 | 1112 | |
| 39097 | 1113 | subsection {* Conditional Entropy *}
 | 
| 1114 | ||
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1115 | definition (in prob_space) | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1116 | "conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y" | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1117 | |
| 40859 | 1118 | abbreviation (in information_space) | 
| 1119 |   conditional_entropy_Pow ("\<H>'(_ | _')") where
 | |
| 36624 | 1120 | "\<H>(X | Y) \<equiv> conditional_entropy b | 
| 43920 | 1121 | \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr> | 
| 1122 | \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y" | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1123 | |
| 40859 | 1124 | lemma (in information_space) conditional_entropy_positive: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1125 | "simple_function M X \<Longrightarrow> simple_function M Y \<Longrightarrow> 0 \<le> \<H>(X | Y)" | 
| 40859 | 1126 | unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive) | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1127 | |
| 40859 | 1128 | lemma (in information_space) conditional_entropy_generic_eq: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1129 |   fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
 | 
| 40859 | 1130 | assumes MX: "finite_random_variable MX X" | 
| 1131 | assumes MZ: "finite_random_variable MZ Z" | |
| 39097 | 1132 | shows "conditional_entropy b MX MZ X Z = | 
| 1133 | - (\<Sum>(x, z)\<in>space MX \<times> space MZ. | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1134 |          joint_distribution X Z {(x, z)} * log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
 | 
| 40859 | 1135 | proof - | 
| 1136 | interpret MX: finite_sigma_algebra MX using MX by simp | |
| 1137 | interpret MZ: finite_sigma_algebra MZ using MZ by simp | |
| 1138 |   let "?XXZ x y z" = "joint_distribution X (\<lambda>x. (X x, Z x)) {(x, y, z)}"
 | |
| 1139 |   let "?XZ x z" = "joint_distribution X Z {(x, z)}"
 | |
| 1140 |   let "?Z z" = "distribution Z {z}"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1141 | let "?f x y z" = "log b (?XXZ x y z * ?Z z / (?XZ x z * ?XZ y z))" | 
| 40859 | 1142 |   { fix x z have "?XXZ x x z = ?XZ x z"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1143 | unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) } | 
| 40859 | 1144 | note this[simp] | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1145 |   { fix x x' :: 'c and z assume "x' \<noteq> x"
 | 
| 40859 | 1146 | then have "?XXZ x x' z = 0" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1147 | by (auto simp: distribution_def empty_measure'[symmetric] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1148 | simp del: empty_measure' intro!: arg_cong[where f=\<mu>']) } | 
| 40859 | 1149 | note this[simp] | 
| 1150 |   { fix x x' z assume *: "x \<in> space MX" "z \<in> space MZ"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1151 | then have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1152 | = (\<Sum>x'\<in>space MX. if x = x' then ?XZ x z * ?f x x z else 0)" | 
| 40859 | 1153 | by (auto intro!: setsum_cong) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1154 | also have "\<dots> = ?XZ x z * ?f x x z" | 
| 40859 | 1155 | using `x \<in> space MX` by (simp add: setsum_cases[OF MX.finite_space]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1156 | also have "\<dots> = ?XZ x z * log b (?Z z / ?XZ x z)" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1157 | also have "\<dots> = - ?XZ x z * log b (?XZ x z / ?Z z)" | 
| 40859 | 1158 | using finite_distribution_order(6)[OF MX MZ] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1159 | by (auto simp: log_simps field_simps zero_less_mult_iff) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1160 | finally have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z) = - ?XZ x z * log b (?XZ x z / ?Z z)" . } | 
| 40859 | 1161 | note * = this | 
| 1162 | show ?thesis | |
| 1163 | unfolding conditional_entropy_def | |
| 1164 | unfolding conditional_mutual_information_generic_eq[OF MX MX MZ] | |
| 1165 | by (auto simp: setsum_cartesian_product' setsum_negf[symmetric] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1166 | setsum_commute[of _ "space MZ"] * | 
| 40859 | 1167 | intro!: setsum_cong) | 
| 39097 | 1168 | qed | 
| 1169 | ||
| 40859 | 1170 | lemma (in information_space) conditional_entropy_eq: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1171 | assumes "simple_function M X" "simple_function M Z" | 
| 40859 | 1172 | shows "\<H>(X | Z) = | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1173 | - (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M. | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1174 |          joint_distribution X Z {(x, z)} *
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1175 |          log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1176 | by (subst conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1177 | simp | 
| 39097 | 1178 | |
| 40859 | 1179 | lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1180 | assumes X: "simple_function M X" and Y: "simple_function M Y" | 
| 40859 | 1181 | shows "\<H>(X | Y) = | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1182 |     -(\<Sum>y\<in>Y`space M. distribution Y {y} *
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1183 |       (\<Sum>x\<in>X`space M. joint_distribution X Y {(x,y)} / distribution Y {(y)} *
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1184 |               log b (joint_distribution X Y {(x,y)} / distribution Y {(y)})))"
 | 
| 40859 | 1185 | unfolding conditional_entropy_eq[OF assms] | 
| 1186 | using finite_distribution_order(5,6)[OF assms[THEN simple_function_imp_finite_random_variable]] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1187 | by (auto simp: setsum_cartesian_product' setsum_commute[of _ "Y`space M"] setsum_right_distrib | 
| 40859 | 1188 | intro!: setsum_cong) | 
| 39097 | 1189 | |
| 40859 | 1190 | lemma (in information_space) conditional_entropy_eq_cartesian_product: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1191 | assumes "simple_function M X" "simple_function M Y" | 
| 40859 | 1192 | shows "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M. | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1193 |     joint_distribution X Y {(x,y)} *
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1194 |     log b (joint_distribution X Y {(x,y)} / distribution Y {y}))"
 | 
| 40859 | 1195 | unfolding conditional_entropy_eq[OF assms] | 
| 1196 | by (auto intro!: setsum_cong simp: setsum_cartesian_product') | |
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1197 | |
| 39097 | 1198 | subsection {* Equalities *}
 | 
| 1199 | ||
| 40859 | 1200 | lemma (in information_space) mutual_information_eq_entropy_conditional_entropy: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1201 | assumes X: "simple_function M X" and Z: "simple_function M Z" | 
| 40859 | 1202 | shows "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" | 
| 1203 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1204 |   let "?XZ x z" = "joint_distribution X Z {(x, z)}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1205 |   let "?Z z" = "distribution Z {z}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1206 |   let "?X x" = "distribution X {x}"
 | 
| 40859 | 1207 | note fX = X[THEN simple_function_imp_finite_random_variable] | 
| 1208 | note fZ = Z[THEN simple_function_imp_finite_random_variable] | |
| 1209 | note finite_distribution_order[OF fX fZ, simp] | |
| 1210 |   { fix x z assume "x \<in> X`space M" "z \<in> Z`space M"
 | |
| 1211 | have "?XZ x z * log b (?XZ x z / (?X x * ?Z z)) = | |
| 1212 | ?XZ x z * log b (?XZ x z / ?Z z) - ?XZ x z * log b (?X x)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1213 | by (auto simp: log_simps zero_le_mult_iff field_simps less_le) } | 
| 40859 | 1214 | note * = this | 
| 1215 | show ?thesis | |
| 1216 | unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1217 | using setsum_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]] | 
| 40859 | 1218 | by (simp add: * setsum_cartesian_product' setsum_subtractf setsum_left_distrib[symmetric] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1219 | setsum_distribution) | 
| 40859 | 1220 | qed | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1221 | |
| 40859 | 1222 | lemma (in information_space) conditional_entropy_less_eq_entropy: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1223 | assumes X: "simple_function M X" and Z: "simple_function M Z" | 
| 40859 | 1224 | shows "\<H>(X | Z) \<le> \<H>(X)" | 
| 36624 | 1225 | proof - | 
| 40859 | 1226 | have "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] . | 
| 1227 | with mutual_information_positive[OF X Z] entropy_positive[OF X] | |
| 36624 | 1228 | show ?thesis by auto | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1229 | qed | 
| 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1230 | |
| 40859 | 1231 | lemma (in information_space) entropy_chain_rule: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1232 | assumes X: "simple_function M X" and Y: "simple_function M Y" | 
| 40859 | 1233 | shows "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)" | 
| 1234 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1235 |   let "?XY x y" = "joint_distribution X Y {(x, y)}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1236 |   let "?Y y" = "distribution Y {y}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1237 |   let "?X x" = "distribution X {x}"
 | 
| 40859 | 1238 | note fX = X[THEN simple_function_imp_finite_random_variable] | 
| 1239 | note fY = Y[THEN simple_function_imp_finite_random_variable] | |
| 1240 | note finite_distribution_order[OF fX fY, simp] | |
| 1241 |   { fix x y assume "x \<in> X`space M" "y \<in> Y`space M"
 | |
| 1242 | have "?XY x y * log b (?XY x y / ?X x) = | |
| 1243 | ?XY x y * log b (?XY x y) - ?XY x y * log b (?X x)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1244 | by (auto simp: log_simps zero_le_mult_iff field_simps less_le) } | 
| 40859 | 1245 | note * = this | 
| 1246 | show ?thesis | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1247 | using setsum_joint_distribution_singleton[OF fY fX] | 
| 40859 | 1248 | unfolding entropy_eq[OF X] conditional_entropy_eq_cartesian_product[OF Y X] entropy_eq_cartesian_product[OF X Y] | 
| 1249 | unfolding joint_distribution_commute_singleton[of Y X] setsum_commute[of _ "X`space M"] | |
| 1250 | by (simp add: * setsum_subtractf setsum_left_distrib[symmetric]) | |
| 1251 | qed | |
| 38656 | 1252 | |
| 39097 | 1253 | section {* Partitioning *}
 | 
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1254 | |
| 36624 | 1255 | definition "subvimage A f g \<longleftrightarrow> (\<forall>x \<in> A. f -` {f x} \<inter> A \<subseteq> g -` {g x} \<inter> A)"
 | 
| 1256 | ||
| 1257 | lemma subvimageI: | |
| 1258 | assumes "\<And>x y. \<lbrakk> x \<in> A ; y \<in> A ; f x = f y \<rbrakk> \<Longrightarrow> g x = g y" | |
| 1259 | shows "subvimage A f g" | |
| 1260 | using assms unfolding subvimage_def by blast | |
| 1261 | ||
| 1262 | lemma subvimageE[consumes 1]: | |
| 1263 | assumes "subvimage A f g" | |
| 1264 | obtains "\<And>x y. \<lbrakk> x \<in> A ; y \<in> A ; f x = f y \<rbrakk> \<Longrightarrow> g x = g y" | |
| 1265 | using assms unfolding subvimage_def by blast | |
| 1266 | ||
| 1267 | lemma subvimageD: | |
| 1268 | "\<lbrakk> subvimage A f g ; x \<in> A ; y \<in> A ; f x = f y \<rbrakk> \<Longrightarrow> g x = g y" | |
| 1269 | using assms unfolding subvimage_def by blast | |
| 1270 | ||
| 1271 | lemma subvimage_subset: | |
| 1272 | "\<lbrakk> subvimage B f g ; A \<subseteq> B \<rbrakk> \<Longrightarrow> subvimage A f g" | |
| 1273 | unfolding subvimage_def by auto | |
| 1274 | ||
| 1275 | lemma subvimage_idem[intro]: "subvimage A g g" | |
| 1276 | by (safe intro!: subvimageI) | |
| 1277 | ||
| 1278 | lemma subvimage_comp_finer[intro]: | |
| 1279 | assumes svi: "subvimage A g h" | |
| 1280 | shows "subvimage A g (f \<circ> h)" | |
| 1281 | proof (rule subvimageI, simp) | |
| 1282 | fix x y assume "x \<in> A" "y \<in> A" "g x = g y" | |
| 1283 | from svi[THEN subvimageD, OF this] | |
| 1284 | show "f (h x) = f (h y)" by simp | |
| 1285 | qed | |
| 1286 | ||
| 1287 | lemma subvimage_comp_gran: | |
| 1288 | assumes svi: "subvimage A g h" | |
| 1289 | assumes inj: "inj_on f (g ` A)" | |
| 1290 | shows "subvimage A (f \<circ> g) h" | |
| 1291 | by (rule subvimageI) (auto intro!: subvimageD[OF svi] simp: inj_on_iff[OF inj]) | |
| 1292 | ||
| 1293 | lemma subvimage_comp: | |
| 1294 | assumes svi: "subvimage (f ` A) g h" | |
| 1295 | shows "subvimage A (g \<circ> f) (h \<circ> f)" | |
| 1296 | by (rule subvimageI) (auto intro!: svi[THEN subvimageD]) | |
| 1297 | ||
| 1298 | lemma subvimage_trans: | |
| 1299 | assumes fg: "subvimage A f g" | |
| 1300 | assumes gh: "subvimage A g h" | |
| 1301 | shows "subvimage A f h" | |
| 1302 | by (rule subvimageI) (auto intro!: fg[THEN subvimageD] gh[THEN subvimageD]) | |
| 1303 | ||
| 1304 | lemma subvimage_translator: | |
| 1305 | assumes svi: "subvimage A f g" | |
| 1306 | shows "\<exists>h. \<forall>x \<in> A. h (f x) = g x" | |
| 1307 | proof (safe intro!: exI[of _ "\<lambda>x. (THE z. z \<in> (g ` (f -` {x} \<inter> A)))"])
 | |
| 1308 | fix x assume "x \<in> A" | |
| 1309 |   show "(THE x'. x' \<in> (g ` (f -` {f x} \<inter> A))) = g x"
 | |
| 1310 | by (rule theI2[of _ "g x"]) | |
| 1311 | (insert `x \<in> A`, auto intro!: svi[THEN subvimageD]) | |
| 1312 | qed | |
| 1313 | ||
| 1314 | lemma subvimage_translator_image: | |
| 1315 | assumes svi: "subvimage A f g" | |
| 1316 | shows "\<exists>h. h ` f ` A = g ` A" | |
| 1317 | proof - | |
| 1318 | from subvimage_translator[OF svi] | |
| 1319 | obtain h where "\<And>x. x \<in> A \<Longrightarrow> h (f x) = g x" by auto | |
| 1320 | thus ?thesis | |
| 1321 | by (auto intro!: exI[of _ h] | |
| 1322 | simp: image_compose[symmetric] comp_def cong: image_cong) | |
| 1323 | qed | |
| 1324 | ||
| 1325 | lemma subvimage_finite: | |
| 1326 | assumes svi: "subvimage A f g" and fin: "finite (f`A)" | |
| 1327 | shows "finite (g`A)" | |
| 1328 | proof - | |
| 1329 | from subvimage_translator_image[OF svi] | |
| 1330 | obtain h where "g`A = h`f`A" by fastsimp | |
| 1331 | with fin show "finite (g`A)" by simp | |
| 1332 | qed | |
| 1333 | ||
| 1334 | lemma subvimage_disj: | |
| 1335 | assumes svi: "subvimage A f g" | |
| 1336 |   shows "f -` {x} \<inter> A \<subseteq> g -` {y} \<inter> A \<or>
 | |
| 1337 |       f -` {x} \<inter> g -` {y} \<inter> A = {}" (is "?sub \<or> ?dist")
 | |
| 1338 | proof (rule disjCI) | |
| 1339 | assume "\<not> ?dist" | |
| 1340 | then obtain z where "z \<in> A" and "x = f z" and "y = g z" by auto | |
| 1341 | thus "?sub" using svi unfolding subvimage_def by auto | |
| 1342 | qed | |
| 1343 | ||
| 1344 | lemma setsum_image_split: | |
| 1345 | assumes svi: "subvimage A f g" and fin: "finite (f ` A)" | |
| 1346 |   shows "(\<Sum>x\<in>f`A. h x) = (\<Sum>y\<in>g`A. \<Sum>x\<in>f`(g -` {y} \<inter> A). h x)"
 | |
| 1347 | (is "?lhs = ?rhs") | |
| 1348 | proof - | |
| 1349 | have "f ` A = | |
| 1350 |       snd ` (SIGMA x : g ` A. f ` (g -` {x} \<inter> A))"
 | |
| 1351 | (is "_ = snd ` ?SIGMA") | |
| 1352 | unfolding image_split_eq_Sigma[symmetric] | |
| 1353 | by (simp add: image_compose[symmetric] comp_def) | |
| 1354 | moreover | |
| 1355 | have snd_inj: "inj_on snd ?SIGMA" | |
| 1356 | unfolding image_split_eq_Sigma[symmetric] | |
| 1357 | by (auto intro!: inj_onI subvimageD[OF svi]) | |
| 1358 | ultimately | |
| 1359 | have "(\<Sum>x\<in>f`A. h x) = (\<Sum>(x,y)\<in>?SIGMA. h y)" | |
| 1360 | by (auto simp: setsum_reindex intro: setsum_cong) | |
| 1361 | also have "... = ?rhs" | |
| 1362 | using subvimage_finite[OF svi fin] fin | |
| 1363 | apply (subst setsum_Sigma[symmetric]) | |
| 1364 | by (auto intro!: finite_subset[of _ "f`A"]) | |
| 1365 | finally show ?thesis . | |
| 1366 | qed | |
| 1367 | ||
| 40859 | 1368 | lemma (in information_space) entropy_partition: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1369 | assumes sf: "simple_function M X" "simple_function M P" | 
| 36624 | 1370 | assumes svi: "subvimage (space M) X P" | 
| 1371 | shows "\<H>(X) = \<H>(P) + \<H>(X|P)" | |
| 1372 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1373 |   let "?XP x p" = "joint_distribution X P {(x, p)}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1374 |   let "?X x" = "distribution X {x}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1375 |   let "?P p" = "distribution P {p}"
 | 
| 40859 | 1376 | note fX = sf(1)[THEN simple_function_imp_finite_random_variable] | 
| 1377 | note fP = sf(2)[THEN simple_function_imp_finite_random_variable] | |
| 1378 | note finite_distribution_order[OF fX fP, simp] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1379 | have "(\<Sum>x\<in>X ` space M. ?X x * log b (?X x)) = | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1380 | (\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M. ?XP x y * log b (?XP x y))" | 
| 36624 | 1381 | proof (subst setsum_image_split[OF svi], | 
| 40859 | 1382 | safe intro!: setsum_mono_zero_cong_left imageI) | 
| 1383 | show "finite (X ` space M)" "finite (X ` space M)" "finite (P ` space M)" | |
| 1384 | using sf unfolding simple_function_def by auto | |
| 1385 | next | |
| 36624 | 1386 | fix p x assume in_space: "p \<in> space M" "x \<in> space M" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1387 | assume "?XP (X x) (P p) * log b (?XP (X x) (P p)) \<noteq> 0" | 
| 36624 | 1388 |     hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M \<noteq> {}" by (auto simp: distribution_def)
 | 
| 1389 | with svi[unfolded subvimage_def, rule_format, OF `x \<in> space M`] | |
| 1390 |     show "x \<in> P -` {P p}" by auto
 | |
| 1391 | next | |
| 1392 | fix p x assume in_space: "p \<in> space M" "x \<in> space M" | |
| 1393 | assume "P x = P p" | |
| 1394 | from this[symmetric] svi[unfolded subvimage_def, rule_format, OF `x \<in> space M`] | |
| 1395 |     have "X -` {X x} \<inter> space M \<subseteq> P -` {P p} \<inter> space M"
 | |
| 1396 | by auto | |
| 1397 |     hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M = X -` {X x} \<inter> space M"
 | |
| 1398 | by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1399 | thus "?X (X x) * log b (?X (X x)) = ?XP (X x) (P p) * log b (?XP (X x) (P p))" | 
| 36624 | 1400 | by (auto simp: distribution_def) | 
| 1401 | qed | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1402 | moreover have "\<And>x y. ?XP x y * log b (?XP x y / ?P y) = | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1403 | ?XP x y * log b (?XP x y) - ?XP x y * log b (?P y)" | 
| 40859 | 1404 | by (auto simp add: log_simps zero_less_mult_iff field_simps) | 
| 1405 | ultimately show ?thesis | |
| 1406 | unfolding sf[THEN entropy_eq] conditional_entropy_eq[OF sf] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1407 | using setsum_joint_distribution_singleton[OF fX fP] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41833diff
changeset | 1408 | by (simp add: setsum_cartesian_product' setsum_subtractf setsum_distribution | 
| 36624 | 1409 | setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"]) | 
| 1410 | qed | |
| 1411 | ||
| 40859 | 1412 | corollary (in information_space) entropy_data_processing: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1413 | assumes X: "simple_function M X" shows "\<H>(f \<circ> X) \<le> \<H>(X)" | 
| 40859 | 1414 | proof - | 
| 1415 | note X | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1416 | moreover have fX: "simple_function M (f \<circ> X)" using X by auto | 
| 40859 | 1417 | moreover have "subvimage (space M) X (f \<circ> X)" by auto | 
| 1418 | ultimately have "\<H>(X) = \<H>(f\<circ>X) + \<H>(X|f\<circ>X)" by (rule entropy_partition) | |
| 1419 | then show "\<H>(f \<circ> X) \<le> \<H>(X)" | |
| 1420 | by (auto intro: conditional_entropy_positive[OF X fX]) | |
| 1421 | qed | |
| 36624 | 1422 | |
| 40859 | 1423 | corollary (in information_space) entropy_of_inj: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1424 | assumes X: "simple_function M X" and inj: "inj_on f (X`space M)" | 
| 36624 | 1425 | shows "\<H>(f \<circ> X) = \<H>(X)" | 
| 1426 | proof (rule antisym) | |
| 40859 | 1427 | show "\<H>(f \<circ> X) \<le> \<H>(X)" using entropy_data_processing[OF X] . | 
| 36624 | 1428 | next | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1429 | have sf: "simple_function M (f \<circ> X)" | 
| 40859 | 1430 | using X by auto | 
| 36624 | 1431 | have "\<H>(X) = \<H>(the_inv_into (X`space M) f \<circ> (f \<circ> X))" | 
| 40859 | 1432 | by (auto intro!: mutual_information_cong simp: entropy_def the_inv_into_f_f[OF inj]) | 
| 36624 | 1433 | also have "... \<le> \<H>(f \<circ> X)" | 
| 40859 | 1434 | using entropy_data_processing[OF sf] . | 
| 36624 | 1435 | finally show "\<H>(X) \<le> \<H>(f \<circ> X)" . | 
| 1436 | qed | |
| 1437 | ||
| 36080 
0d9affa4e73c
Added Information theory and Example: dining cryptographers
 hoelzl parents: diff
changeset | 1438 | end |