author | wenzelm |
Mon, 20 May 2024 15:43:51 +0200 | |
changeset 80182 | 29f2b8ff84f3 |
parent 80034 | 95b4fb2b5359 |
child 80521 | 5c691b178e08 |
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 |
||
61808 | 6 |
section \<open>Information theory\<close> |
42067 | 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:
41095
diff
changeset
|
9 |
imports |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
10 |
Independent_Family |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
11 |
begin |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
12 |
|
56994 | 13 |
subsection "Information theory" |
38656 | 14 |
|
40859 | 15 |
locale information_space = prob_space + |
38656 | 16 |
fixes b :: real assumes b_gt_1: "1 < b" |
17 |
||
40859 | 18 |
context information_space |
38656 | 19 |
begin |
20 |
||
69597 | 21 |
text \<open>Introduce some simplification rules for logarithm of base \<^term>\<open>b\<close>.\<close> |
40859 | 22 |
|
23 |
lemma log_neg_const: |
|
24 |
assumes "x \<le> 0" |
|
25 |
shows "log b x = log b 0" |
|
36624 | 26 |
proof - |
40859 | 27 |
{ fix u :: real |
28 |
have "x \<le> 0" by fact |
|
29 |
also have "0 < exp u" |
|
30 |
using exp_gt_zero . |
|
31 |
finally have "exp u \<noteq> x" |
|
32 |
by auto } |
|
33 |
then show "log b x = log b 0" |
|
60017
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents:
58876
diff
changeset
|
34 |
by (simp add: log_def ln_real_def) |
38656 | 35 |
qed |
36 |
||
40859 | 37 |
lemma log_mult_eq: |
38 |
"log b (A * B) = (if 0 < A * B then log b \<bar>A\<bar> + log b \<bar>B\<bar> else log b 0)" |
|
79772
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
39 |
using log_mult[of "\<bar>A\<bar>" "\<bar>B\<bar>"] b_gt_1 log_neg_const[of "A * B"] |
40859 | 40 |
by (auto simp: zero_less_mult_iff mult_le_0_iff) |
38656 | 41 |
|
40859 | 42 |
lemma log_inverse_eq: |
43 |
"log b (inverse B) = (if 0 < B then - log b B else log b 0)" |
|
79772
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
44 |
using ln_inverse log_def log_neg_const by force |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
45 |
|
40859 | 46 |
lemma log_divide_eq: |
79772
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
47 |
"log b (A / B) = (if 0 < A * B then (log b \<bar>A\<bar>) - log b \<bar>B\<bar> else log b 0)" |
40859 | 48 |
unfolding divide_inverse log_mult_eq log_inverse_eq abs_inverse |
49 |
by (auto simp: zero_less_mult_iff mult_le_0_iff) |
|
38656 | 50 |
|
40859 | 51 |
lemmas log_simps = log_mult_eq log_inverse_eq log_divide_eq |
38656 | 52 |
|
53 |
end |
|
54 |
||
39097 | 55 |
subsection "Kullback$-$Leibler divergence" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
56 |
|
61808 | 57 |
text \<open>The Kullback$-$Leibler divergence is also known as relative entropy or |
58 |
Kullback$-$Leibler distance.\<close> |
|
39097 | 59 |
|
60 |
definition |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
61 |
"entropy_density b M N = log b \<circ> enn2real \<circ> RN_deriv M N" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
62 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
63 |
definition |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
64 |
"KL_divergence b M N = integral\<^sup>L N (entropy_density b M N)" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
65 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
66 |
lemma measurable_entropy_density[measurable]: "entropy_density b M N \<in> borel_measurable M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
67 |
unfolding entropy_density_def by auto |
50003 | 68 |
|
47694 | 69 |
lemma (in sigma_finite_measure) KL_density: |
70 |
fixes f :: "'a \<Rightarrow> real" |
|
71 |
assumes "1 < b" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
72 |
assumes f[measurable]: "f \<in> borel_measurable M" and nn: "AE x in M. 0 \<le> f x" |
47694 | 73 |
shows "KL_divergence b M (density M f) = (\<integral>x. f x * log b (f x) \<partial>M)" |
74 |
unfolding KL_divergence_def |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
75 |
proof (subst integral_real_density) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
76 |
show [measurable]: "entropy_density b M (density M (\<lambda>x. ennreal (f x))) \<in> borel_measurable M" |
49776 | 77 |
using f |
50003 | 78 |
by (auto simp: comp_def entropy_density_def) |
47694 | 79 |
have "density M (RN_deriv M (density M f)) = density M f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
80 |
using f nn by (intro density_RN_deriv_density) auto |
47694 | 81 |
then have eq: "AE x in M. RN_deriv M (density M f) x = f x" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
82 |
using f nn by (intro density_unique) auto |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
83 |
have "AE x in M. f x * entropy_density b M (density M (\<lambda>x. ennreal (f x))) x = |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
84 |
f x * log b (f x)" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
85 |
using eq nn by (auto simp: entropy_density_def) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
86 |
then show "(\<integral>x. f x * entropy_density b M (density M (\<lambda>x. ennreal (f x))) x \<partial>M) = (\<integral>x. f x * log b (f x) \<partial>M)" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
87 |
by (intro integral_cong_AE) measurable |
47694 | 88 |
qed fact+ |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
89 |
|
47694 | 90 |
lemma (in sigma_finite_measure) KL_density_density: |
91 |
fixes f g :: "'a \<Rightarrow> real" |
|
92 |
assumes "1 < b" |
|
93 |
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" |
|
94 |
assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" |
|
95 |
assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0" |
|
96 |
shows "KL_divergence b (density M f) (density M g) = (\<integral>x. g x * log b (g x / f x) \<partial>M)" |
|
97 |
proof - |
|
98 |
interpret Mf: sigma_finite_measure "density M f" |
|
99 |
using f by (subst sigma_finite_iff_density_finite) auto |
|
100 |
have "KL_divergence b (density M f) (density M g) = |
|
101 |
KL_divergence b (density M f) (density (density M f) (\<lambda>x. g x / f x))" |
|
102 |
using f g ac by (subst density_density_divide) simp_all |
|
103 |
also have "\<dots> = (\<integral>x. (g x / f x) * log b (g x / f x) \<partial>density M f)" |
|
61808 | 104 |
using f g \<open>1 < b\<close> by (intro Mf.KL_density) (auto simp: AE_density) |
47694 | 105 |
also have "\<dots> = (\<integral>x. g x * log b (g x / f x) \<partial>M)" |
61808 | 106 |
using ac f g \<open>1 < b\<close> by (subst integral_density) (auto intro!: integral_cong_AE) |
47694 | 107 |
finally show ?thesis . |
108 |
qed |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
109 |
|
47694 | 110 |
lemma (in information_space) KL_gt_0: |
111 |
fixes D :: "'a \<Rightarrow> real" |
|
112 |
assumes "prob_space (density M D)" |
|
113 |
assumes D: "D \<in> borel_measurable M" "AE x in M. 0 \<le> D x" |
|
114 |
assumes int: "integrable M (\<lambda>x. D x * log b (D x))" |
|
115 |
assumes A: "density M D \<noteq> M" |
|
116 |
shows "0 < KL_divergence b M (density M D)" |
|
117 |
proof - |
|
118 |
interpret N: prob_space "density M D" by fact |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
119 |
|
47694 | 120 |
obtain A where "A \<in> sets M" "emeasure (density M D) A \<noteq> emeasure M A" |
61808 | 121 |
using measure_eqI[of "density M D" M] \<open>density M D \<noteq> M\<close> by auto |
47694 | 122 |
|
123 |
let ?D_set = "{x\<in>space M. D x \<noteq> 0}" |
|
124 |
have [simp, intro]: "?D_set \<in> sets M" |
|
125 |
using D by auto |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
126 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
127 |
have D_neg: "(\<integral>\<^sup>+ x. ennreal (- D x) \<partial>M) = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
128 |
using D by (subst nn_integral_0_iff_AE) (auto simp: ennreal_neg) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
129 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
130 |
have "(\<integral>\<^sup>+ x. ennreal (D x) \<partial>M) = emeasure (density M D) (space M)" |
56996 | 131 |
using D by (simp add: emeasure_density cong: nn_integral_cong) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
132 |
then have D_pos: "(\<integral>\<^sup>+ x. ennreal (D x) \<partial>M) = 1" |
47694 | 133 |
using N.emeasure_space_1 by simp |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
134 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
135 |
have "integrable M D" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
136 |
using D D_pos D_neg unfolding real_integrable_def real_lebesgue_integral_def by simp_all |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
137 |
then have "integral\<^sup>L M D = 1" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
138 |
using D D_pos D_neg by (simp add: real_lebesgue_integral_def) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
139 |
|
47694 | 140 |
have "0 \<le> 1 - measure M ?D_set" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
141 |
using prob_le_1 by (auto simp: field_simps) |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
142 |
also have "\<dots> = (\<integral> x. D x - indicator ?D_set x \<partial>M)" |
61808 | 143 |
using \<open>integrable M D\<close> \<open>integral\<^sup>L M D = 1\<close> |
47694 | 144 |
by (simp add: emeasure_eq_measure) |
145 |
also have "\<dots> < (\<integral> x. D x * (ln b * log b (D x)) \<partial>M)" |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
146 |
proof (rule integral_less_AE) |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
147 |
show "integrable M (\<lambda>x. D x - indicator ?D_set x)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
148 |
using \<open>integrable M D\<close> by (auto simp: less_top[symmetric]) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
149 |
next |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
150 |
from integrable_mult_left(1)[OF int, of "ln b"] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
151 |
show "integrable M (\<lambda>x. D x * (ln b * log b (D x)))" |
47694 | 152 |
by (simp add: ac_simps) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
153 |
next |
47694 | 154 |
show "emeasure M {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<noteq> 0" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
155 |
proof |
47694 | 156 |
assume eq_0: "emeasure M {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} = 0" |
157 |
then have disj: "AE x in M. D x = 1 \<or> D x = 0" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50003
diff
changeset
|
158 |
using D(1) by (auto intro!: AE_I[OF subset_refl] sets.sets_Collect) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
159 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
160 |
have "emeasure M {x\<in>space M. D x = 1} = (\<integral>\<^sup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
161 |
using D(1) by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
162 |
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (D x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
163 |
using disj by (auto intro!: nn_integral_cong_AE simp: indicator_def one_ennreal_def) |
47694 | 164 |
finally have "AE x in M. D x = 1" |
165 |
using D D_pos by (intro AE_I_eq_1) auto |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
166 |
then have "(\<integral>\<^sup>+x. indicator A x\<partial>M) = (\<integral>\<^sup>+x. ennreal (D x) * indicator A x\<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
167 |
by (intro nn_integral_cong_AE) (auto simp: one_ennreal_def[symmetric]) |
47694 | 168 |
also have "\<dots> = density M D A" |
61808 | 169 |
using \<open>A \<in> sets M\<close> D by (simp add: emeasure_density) |
170 |
finally show False using \<open>A \<in> sets M\<close> \<open>emeasure (density M D) A \<noteq> emeasure M A\<close> by simp |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
171 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
172 |
show "{x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<in> sets M" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50003
diff
changeset
|
173 |
using D(1) by (auto intro: sets.sets_Collect_conj) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
174 |
|
47694 | 175 |
show "AE t in M. t \<in> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<longrightarrow> |
176 |
D t - indicator ?D_set t \<noteq> D t * (ln b * log b (D t))" |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
177 |
using D(2) |
47694 | 178 |
proof (eventually_elim, safe) |
179 |
fix t assume Dt: "t \<in> space M" "D t \<noteq> 1" "D t \<noteq> 0" "0 \<le> D t" |
|
180 |
and eq: "D t - indicator ?D_set t = D t * (ln b * log b (D t))" |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
181 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
182 |
have "D t - 1 = D t - indicator ?D_set t" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
183 |
using Dt by simp |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
184 |
also note eq |
47694 | 185 |
also have "D t * (ln b * log b (D t)) = - D t * ln (1 / D t)" |
61808 | 186 |
using b_gt_1 \<open>D t \<noteq> 0\<close> \<open>0 \<le> D t\<close> |
47694 | 187 |
by (simp add: log_def ln_div less_le) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
188 |
finally have "ln (1 / D t) = 1 / D t - 1" |
61808 | 189 |
using \<open>D t \<noteq> 0\<close> by (auto simp: field_simps) |
190 |
from ln_eq_minus_one[OF _ this] \<open>D t \<noteq> 0\<close> \<open>0 \<le> D t\<close> \<open>D t \<noteq> 1\<close> |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
191 |
show False by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
192 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
193 |
|
47694 | 194 |
show "AE t in M. D t - indicator ?D_set t \<le> D t * (ln b * log b (D t))" |
195 |
using D(2) AE_space |
|
196 |
proof eventually_elim |
|
197 |
fix t assume "t \<in> space M" "0 \<le> D t" |
|
198 |
show "D t - indicator ?D_set t \<le> D t * (ln b * log b (D t))" |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
199 |
proof cases |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
200 |
assume asm: "D t \<noteq> 0" |
61808 | 201 |
then have "0 < D t" using \<open>0 \<le> D t\<close> by auto |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
202 |
then have "0 < 1 / D t" by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
203 |
have "D t - indicator ?D_set t \<le> - D t * (1 / D t - 1)" |
61808 | 204 |
using asm \<open>t \<in> space M\<close> by (simp add: field_simps) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
205 |
also have "- D t * (1 / D t - 1) \<le> - D t * ln (1 / D t)" |
61808 | 206 |
using ln_le_minus_one \<open>0 < 1 / D t\<close> by (intro mult_left_mono_neg) auto |
47694 | 207 |
also have "\<dots> = D t * (ln b * log b (D t))" |
61808 | 208 |
using \<open>0 < D t\<close> b_gt_1 |
47694 | 209 |
by (simp_all add: log_def ln_div) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
210 |
finally show ?thesis by simp |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
211 |
qed simp |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
212 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
213 |
qed |
47694 | 214 |
also have "\<dots> = (\<integral> x. ln b * (D x * log b (D x)) \<partial>M)" |
215 |
by (simp add: ac_simps) |
|
216 |
also have "\<dots> = ln b * (\<integral> x. D x * log b (D x) \<partial>M)" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
217 |
using int by simp |
47694 | 218 |
finally show ?thesis |
219 |
using b_gt_1 D by (subst KL_density) (auto simp: zero_less_mult_iff) |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
220 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
221 |
|
47694 | 222 |
lemma (in sigma_finite_measure) KL_same_eq_0: "KL_divergence b M M = 0" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
223 |
proof - |
47694 | 224 |
have "AE x in M. 1 = RN_deriv M M x" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
225 |
proof (rule RN_deriv_unique) |
47694 | 226 |
show "density M (\<lambda>x. 1) = M" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
227 |
by (simp add: density_1) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
228 |
qed auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
229 |
then have "AE x in M. log b (enn2real (RN_deriv M M x)) = 0" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
230 |
by (elim AE_mp) simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
231 |
from integral_cong_AE[OF _ _ this] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
232 |
have "integral\<^sup>L M (entropy_density b M M) = 0" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
233 |
by (simp add: entropy_density_def comp_def) |
47694 | 234 |
then show "KL_divergence b M M = 0" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
235 |
unfolding KL_divergence_def |
47694 | 236 |
by auto |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
237 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
238 |
|
47694 | 239 |
lemma (in information_space) KL_eq_0_iff_eq: |
240 |
fixes D :: "'a \<Rightarrow> real" |
|
241 |
assumes "prob_space (density M D)" |
|
242 |
assumes D: "D \<in> borel_measurable M" "AE x in M. 0 \<le> D x" |
|
243 |
assumes int: "integrable M (\<lambda>x. D x * log b (D x))" |
|
244 |
shows "KL_divergence b M (density M D) = 0 \<longleftrightarrow> density M D = M" |
|
245 |
using KL_same_eq_0[of b] KL_gt_0[OF assms] |
|
246 |
by (auto simp: less_le) |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
247 |
|
47694 | 248 |
lemma (in information_space) KL_eq_0_iff_eq_ac: |
249 |
fixes D :: "'a \<Rightarrow> real" |
|
250 |
assumes "prob_space N" |
|
251 |
assumes ac: "absolutely_continuous M N" "sets N = sets M" |
|
252 |
assumes int: "integrable N (entropy_density b M N)" |
|
253 |
shows "KL_divergence b M N = 0 \<longleftrightarrow> N = M" |
|
41833
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
hoelzl
parents:
41689
diff
changeset
|
254 |
proof - |
47694 | 255 |
interpret N: prob_space N by fact |
256 |
have "finite_measure N" by unfold_locales |
|
74362 | 257 |
from real_RN_deriv[OF this ac] obtain D |
258 |
where D: |
|
259 |
"random_variable borel D" |
|
260 |
"AE x in M. RN_deriv M N x = ennreal (D x)" |
|
261 |
"AE x in N. 0 < D x" |
|
262 |
"\<And>x. 0 \<le> D x" |
|
263 |
by this auto |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
264 |
|
47694 | 265 |
have "N = density M (RN_deriv M N)" |
266 |
using ac by (rule density_RN_deriv[symmetric]) |
|
267 |
also have "\<dots> = density M D" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
268 |
using D by (auto intro!: density_cong) |
47694 | 269 |
finally have N: "N = density M D" . |
41833
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
hoelzl
parents:
41689
diff
changeset
|
270 |
|
47694 | 271 |
from absolutely_continuous_AE[OF ac(2,1) D(2)] D b_gt_1 ac measurable_entropy_density |
272 |
have "integrable N (\<lambda>x. log b (D x))" |
|
273 |
by (intro integrable_cong_AE[THEN iffD2, OF _ _ _ int]) |
|
274 |
(auto simp: N entropy_density_def) |
|
275 |
with D b_gt_1 have "integrable M (\<lambda>x. D x * log b (D x))" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
276 |
by (subst integrable_real_density[symmetric]) (auto simp: N[symmetric] comp_def) |
61808 | 277 |
with \<open>prob_space N\<close> D show ?thesis |
47694 | 278 |
unfolding N |
279 |
by (intro KL_eq_0_iff_eq) auto |
|
41833
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
hoelzl
parents:
41689
diff
changeset
|
280 |
qed |
563bea92b2c0
add lemma KL_divergence_vimage, mutual_information_generic
hoelzl
parents:
41689
diff
changeset
|
281 |
|
47694 | 282 |
lemma (in information_space) KL_nonneg: |
283 |
assumes "prob_space (density M D)" |
|
284 |
assumes D: "D \<in> borel_measurable M" "AE x in M. 0 \<le> D x" |
|
285 |
assumes int: "integrable M (\<lambda>x. D x * log b (D x))" |
|
286 |
shows "0 \<le> KL_divergence b M (density M D)" |
|
287 |
using KL_gt_0[OF assms] by (cases "density M D = M") (auto simp: KL_same_eq_0) |
|
40859 | 288 |
|
47694 | 289 |
lemma (in sigma_finite_measure) KL_density_density_nonneg: |
290 |
fixes f g :: "'a \<Rightarrow> real" |
|
291 |
assumes "1 < b" |
|
292 |
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "prob_space (density M f)" |
|
293 |
assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" "prob_space (density M g)" |
|
294 |
assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0" |
|
295 |
assumes int: "integrable M (\<lambda>x. g x * log b (g x / f x))" |
|
296 |
shows "0 \<le> KL_divergence b (density M f) (density M g)" |
|
297 |
proof - |
|
298 |
interpret Mf: prob_space "density M f" by fact |
|
61169 | 299 |
interpret Mf: information_space "density M f" b by standard fact |
47694 | 300 |
have eq: "density (density M f) (\<lambda>x. g x / f x) = density M g" (is "?DD = _") |
301 |
using f g ac by (subst density_density_divide) simp_all |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
302 |
|
47694 | 303 |
have "0 \<le> KL_divergence b (density M f) (density (density M f) (\<lambda>x. g x / f x))" |
304 |
proof (rule Mf.KL_nonneg) |
|
305 |
show "prob_space ?DD" unfolding eq by fact |
|
306 |
from f g show "(\<lambda>x. g x / f x) \<in> borel_measurable (density M f)" |
|
307 |
by auto |
|
308 |
show "AE x in density M f. 0 \<le> g x / f x" |
|
56571
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents:
56544
diff
changeset
|
309 |
using f g by (auto simp: AE_density) |
47694 | 310 |
show "integrable (density M f) (\<lambda>x. g x / f x * log b (g x / f x))" |
61808 | 311 |
using \<open>1 < b\<close> f g ac |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
312 |
by (subst integrable_density) |
47694 | 313 |
(auto intro!: integrable_cong_AE[THEN iffD2, OF _ _ _ int] measurable_If) |
314 |
qed |
|
315 |
also have "\<dots> = KL_divergence b (density M f) (density M g)" |
|
316 |
using f g ac by (subst density_density_divide) simp_all |
|
317 |
finally show ?thesis . |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
318 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
319 |
|
61808 | 320 |
subsection \<open>Finite Entropy\<close> |
49803 | 321 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
322 |
definition (in information_space) finite_entropy :: "'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> real) \<Rightarrow> bool" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
323 |
where |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
324 |
"finite_entropy S X f \<longleftrightarrow> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
325 |
distributed M S X f \<and> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
326 |
integrable S (\<lambda>x. f x * log b (f x)) \<and> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
327 |
(\<forall>x\<in>space S. 0 \<le> f x)" |
49803 | 328 |
|
329 |
lemma (in information_space) finite_entropy_simple_function: |
|
330 |
assumes X: "simple_function M X" |
|
331 |
shows "finite_entropy (count_space (X`space M)) X (\<lambda>a. measure M {x \<in> space M. X x = a})" |
|
332 |
unfolding finite_entropy_def |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
333 |
proof safe |
49803 | 334 |
have [simp]: "finite (X ` space M)" |
335 |
using X by (auto simp: simple_function_def) |
|
336 |
then show "integrable (count_space (X ` space M)) |
|
337 |
(\<lambda>x. prob {xa \<in> space M. X xa = x} * log b (prob {xa \<in> space M. X xa = x}))" |
|
338 |
by (rule integrable_count_space) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
339 |
have d: "distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (if x \<in> X`space M then prob {xa \<in> space M. X xa = x} else 0))" |
49803 | 340 |
by (rule distributed_simple_function_superset[OF X]) (auto intro!: arg_cong[where f=prob]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
341 |
show "distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (prob {xa \<in> space M. X xa = x}))" |
49803 | 342 |
by (rule distributed_cong_density[THEN iffD1, OF _ _ _ d]) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
343 |
qed (rule measure_nonneg) |
49803 | 344 |
|
345 |
lemma ac_fst: |
|
346 |
assumes "sigma_finite_measure T" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
347 |
shows "absolutely_continuous S (distr (S \<Otimes>\<^sub>M T) S fst)" |
49803 | 348 |
proof - |
349 |
interpret sigma_finite_measure T by fact |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
350 |
{ fix A assume A: "A \<in> sets S" "emeasure S A = 0" |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
351 |
then have "fst -` A \<inter> space (S \<Otimes>\<^sub>M T) = A \<times> space T" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50003
diff
changeset
|
352 |
by (auto simp: space_pair_measure dest!: sets.sets_into_space) |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
353 |
with A have "emeasure (S \<Otimes>\<^sub>M T) (fst -` A \<inter> space (S \<Otimes>\<^sub>M T)) = 0" |
49803 | 354 |
by (simp add: emeasure_pair_measure_Times) } |
355 |
then show ?thesis |
|
356 |
unfolding absolutely_continuous_def |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
357 |
by (metis emeasure_distr measurable_fst null_setsD1 null_setsD2 null_setsI sets_distr subsetI) |
49803 | 358 |
qed |
359 |
||
360 |
lemma ac_snd: |
|
361 |
assumes "sigma_finite_measure T" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
362 |
shows "absolutely_continuous T (distr (S \<Otimes>\<^sub>M T) T snd)" |
49803 | 363 |
proof - |
364 |
interpret sigma_finite_measure T by fact |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
365 |
{ fix A assume A: "A \<in> sets T" "emeasure T A = 0" |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
366 |
then have "snd -` A \<inter> space (S \<Otimes>\<^sub>M T) = space S \<times> A" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50003
diff
changeset
|
367 |
by (auto simp: space_pair_measure dest!: sets.sets_into_space) |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
368 |
with A have "emeasure (S \<Otimes>\<^sub>M T) (snd -` A \<inter> space (S \<Otimes>\<^sub>M T)) = 0" |
49803 | 369 |
by (simp add: emeasure_pair_measure_Times) } |
370 |
then show ?thesis |
|
371 |
unfolding absolutely_continuous_def |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
372 |
by (metis emeasure_distr measurable_snd null_setsD1 null_setsD2 null_setsI sets_distr subsetI) |
49803 | 373 |
qed |
374 |
||
375 |
lemma (in information_space) finite_entropy_integrable: |
|
376 |
"finite_entropy S X Px \<Longrightarrow> integrable S (\<lambda>x. Px x * log b (Px x))" |
|
377 |
unfolding finite_entropy_def by auto |
|
378 |
||
379 |
lemma (in information_space) finite_entropy_distributed: |
|
380 |
"finite_entropy S X Px \<Longrightarrow> distributed M S X Px" |
|
381 |
unfolding finite_entropy_def by auto |
|
382 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
383 |
lemma (in information_space) finite_entropy_nn: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
384 |
"finite_entropy S X Px \<Longrightarrow> x \<in> space S \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
385 |
by (auto simp: finite_entropy_def) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
386 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
387 |
lemma (in information_space) finite_entropy_measurable: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
388 |
"finite_entropy S X Px \<Longrightarrow> Px \<in> S \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
389 |
using distributed_real_measurable[of S Px M X] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
390 |
finite_entropy_nn[of S X Px] finite_entropy_distributed[of S X Px] by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
391 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
392 |
lemma (in information_space) subdensity_finite_entropy: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
393 |
fixes g :: "'b \<Rightarrow> real" and f :: "'c \<Rightarrow> real" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
394 |
assumes T: "T \<in> measurable P Q" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
395 |
assumes f: "finite_entropy P X f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
396 |
assumes g: "finite_entropy Q Y g" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
397 |
assumes Y: "Y = T \<circ> X" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
398 |
shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
399 |
using subdensity[OF T, of M X "\<lambda>x. ennreal (f x)" Y "\<lambda>x. ennreal (g x)"] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
400 |
finite_entropy_distributed[OF f] finite_entropy_distributed[OF g] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
401 |
finite_entropy_nn[OF f] finite_entropy_nn[OF g] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
402 |
assms |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
403 |
by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
404 |
|
49803 | 405 |
lemma (in information_space) finite_entropy_integrable_transform: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
406 |
"finite_entropy S X Px \<Longrightarrow> distributed M T Y Py \<Longrightarrow> (\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x) \<Longrightarrow> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
407 |
X = (\<lambda>x. f (Y x)) \<Longrightarrow> f \<in> measurable T S \<Longrightarrow> integrable T (\<lambda>x. Py x * log b (Px (f x)))" |
49803 | 408 |
using distributed_transform_integrable[of M T Y Py S X Px f "\<lambda>x. log b (Px x)"] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
409 |
using distributed_real_measurable[of S Px M X] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
410 |
by (auto simp: finite_entropy_def) |
49803 | 411 |
|
61808 | 412 |
subsection \<open>Mutual Information\<close> |
39097 | 413 |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
414 |
definition (in prob_space) |
38656 | 415 |
"mutual_information b S T X Y = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
416 |
KL_divergence b (distr M S X \<Otimes>\<^sub>M distr M T Y) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)))" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
417 |
|
47694 | 418 |
lemma (in information_space) mutual_information_indep_vars: |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
419 |
fixes S T X Y |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
420 |
defines "P \<equiv> distr M S X \<Otimes>\<^sub>M distr M T Y" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
421 |
defines "Q \<equiv> distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
422 |
shows "indep_var S X T Y \<longleftrightarrow> |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
423 |
(random_variable S X \<and> random_variable T Y \<and> |
47694 | 424 |
absolutely_continuous P Q \<and> integrable Q (entropy_density b P Q) \<and> |
425 |
mutual_information b S T X Y = 0)" |
|
426 |
unfolding indep_var_distribution_eq |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
427 |
proof safe |
50003 | 428 |
assume rv[measurable]: "random_variable S X" "random_variable T Y" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
429 |
|
47694 | 430 |
interpret X: prob_space "distr M S X" |
431 |
by (rule prob_space_distr) fact |
|
432 |
interpret Y: prob_space "distr M T Y" |
|
433 |
by (rule prob_space_distr) fact |
|
61169 | 434 |
interpret XY: pair_prob_space "distr M S X" "distr M T Y" by standard |
435 |
interpret P: information_space P b unfolding P_def by standard (rule b_gt_1) |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
436 |
|
47694 | 437 |
interpret Q: prob_space Q unfolding Q_def |
50003 | 438 |
by (rule prob_space_distr) simp |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
439 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
440 |
{ assume "distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
47694 | 441 |
then have [simp]: "Q = P" unfolding Q_def P_def by simp |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
442 |
|
47694 | 443 |
show ac: "absolutely_continuous P Q" by (simp add: absolutely_continuous_def) |
444 |
then have ed: "entropy_density b P Q \<in> borel_measurable P" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
445 |
by simp |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
446 |
|
47694 | 447 |
have "AE x in P. 1 = RN_deriv P Q x" |
448 |
proof (rule P.RN_deriv_unique) |
|
449 |
show "density P (\<lambda>x. 1) = Q" |
|
61808 | 450 |
unfolding \<open>Q = P\<close> by (intro measure_eqI) (auto simp: emeasure_density) |
47694 | 451 |
qed auto |
452 |
then have ae_0: "AE x in P. entropy_density b P Q x = 0" |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
453 |
by (auto simp: entropy_density_def) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
454 |
then have "integrable P (entropy_density b P Q) \<longleftrightarrow> integrable Q (\<lambda>x. 0::real)" |
61808 | 455 |
using ed unfolding \<open>Q = P\<close> by (intro integrable_cong_AE) auto |
47694 | 456 |
then show "integrable Q (entropy_density b P Q)" by simp |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
457 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
458 |
from ae_0 have "mutual_information b S T X Y = (\<integral>x. 0 \<partial>P)" |
61808 | 459 |
unfolding mutual_information_def KL_divergence_def P_def[symmetric] Q_def[symmetric] \<open>Q = P\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
460 |
by (intro integral_cong_AE) auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
461 |
then show "mutual_information b S T X Y = 0" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
462 |
by simp } |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
463 |
|
47694 | 464 |
{ assume ac: "absolutely_continuous P Q" |
465 |
assume int: "integrable Q (entropy_density b P Q)" |
|
466 |
assume I_eq_0: "mutual_information b S T X Y = 0" |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
467 |
|
47694 | 468 |
have eq: "Q = P" |
469 |
proof (rule P.KL_eq_0_iff_eq_ac[THEN iffD1]) |
|
470 |
show "prob_space Q" by unfold_locales |
|
471 |
show "absolutely_continuous P Q" by fact |
|
472 |
show "integrable Q (entropy_density b P Q)" by fact |
|
473 |
show "sets Q = sets P" by (simp add: P_def Q_def sets_pair_measure) |
|
474 |
show "KL_divergence b P Q = 0" |
|
475 |
using I_eq_0 unfolding mutual_information_def by (simp add: P_def Q_def) |
|
476 |
qed |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
477 |
then show "distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
47694 | 478 |
unfolding P_def Q_def .. } |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
479 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42148
diff
changeset
|
480 |
|
40859 | 481 |
abbreviation (in information_space) |
482 |
mutual_information_Pow ("\<I>'(_ ; _')") where |
|
47694 | 483 |
"\<I>(X ; Y) \<equiv> mutual_information b (count_space (X`space M)) (count_space (Y`space M)) 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:
41661
diff
changeset
|
484 |
|
47694 | 485 |
lemma (in information_space) |
486 |
fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" |
|
49803 | 487 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
488 |
assumes Fx: "finite_entropy S X Px" and Fy: "finite_entropy T Y Py" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
489 |
assumes Fxy: "finite_entropy (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
49803 | 490 |
defines "f \<equiv> \<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
491 |
shows mutual_information_distr': "mutual_information b S T X Y = integral\<^sup>L (S \<Otimes>\<^sub>M T) f" (is "?M = ?R") |
49803 | 492 |
and mutual_information_nonneg': "0 \<le> mutual_information b S T X Y" |
493 |
proof - |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
494 |
have Px: "distributed M S X Px" and Px_nn: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x" |
49803 | 495 |
using Fx by (auto simp: finite_entropy_def) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
496 |
have Py: "distributed M T Y Py" and Py_nn: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x" |
49803 | 497 |
using Fy by (auto simp: finite_entropy_def) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
498 |
have Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
499 |
and Pxy_nn: "\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
500 |
"\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
501 |
using Fxy by (auto simp: finite_entropy_def space_pair_measure) |
49803 | 502 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
503 |
have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
504 |
using Px Px_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
505 |
have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
506 |
using Py Py_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
507 |
have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
508 |
using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
509 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
510 |
have X[measurable]: "random_variable S X" |
50003 | 511 |
using Px by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
512 |
have Y[measurable]: "random_variable T Y" |
50003 | 513 |
using Py by auto |
49803 | 514 |
interpret S: sigma_finite_measure S by fact |
515 |
interpret T: sigma_finite_measure T by fact |
|
516 |
interpret ST: pair_sigma_finite S T .. |
|
517 |
interpret X: prob_space "distr M S X" using X by (rule prob_space_distr) |
|
518 |
interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr) |
|
519 |
interpret XY: pair_prob_space "distr M S X" "distr M T Y" .. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
520 |
let ?P = "S \<Otimes>\<^sub>M T" |
49803 | 521 |
let ?D = "distr M ?P (\<lambda>x. (X x, Y x))" |
522 |
||
523 |
{ fix A assume "A \<in> sets S" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
524 |
with X[THEN measurable_space] Y[THEN measurable_space] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
525 |
have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
526 |
by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } |
49803 | 527 |
note marginal_eq1 = this |
528 |
{ fix A assume "A \<in> sets T" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
529 |
with X[THEN measurable_space] Y[THEN measurable_space] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
530 |
have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
531 |
by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } |
49803 | 532 |
note marginal_eq2 = this |
533 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
534 |
have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
535 |
unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] |
49803 | 536 |
proof (subst pair_measure_density) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
537 |
show "(\<lambda>x. ennreal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ennreal (Py y)) \<in> borel_measurable T" |
49803 | 538 |
using Px Py by (auto simp: distributed_def) |
539 |
show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] .. |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
540 |
show "density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). ennreal (Px x) * ennreal (Py y)) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
541 |
density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
542 |
using Px_nn Py_nn by (auto intro!: density_cong simp: distributed_def ennreal_mult space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
543 |
qed fact |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
544 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
545 |
have M: "?M = KL_divergence b (density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ennreal (Pxy x)))" |
49803 | 546 |
unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] .. |
547 |
||
548 |
from Px Py have f: "(\<lambda>x. Px (fst x) * Py (snd x)) \<in> borel_measurable ?P" |
|
549 |
by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'') |
|
550 |
have PxPy_nonneg: "AE x in ?P. 0 \<le> Px (fst x) * Py (snd x)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
551 |
using Px_nn Py_nn by (auto simp: space_pair_measure) |
49803 | 552 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
553 |
have A: "(AE x in ?P. Px (fst x) = 0 \<longrightarrow> Pxy x = 0)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
554 |
by (rule subdensity_real[OF measurable_fst Pxy Px]) (insert Px_nn Pxy_nn, auto simp: space_pair_measure) |
49803 | 555 |
moreover |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
556 |
have B: "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
557 |
by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure) |
49803 | 558 |
ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
559 |
by auto |
49803 | 560 |
|
561 |
show "?M = ?R" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
562 |
unfolding M f_def using Pxy_nn Px_nn Py_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
563 |
by (intro ST.KL_density_density b_gt_1 f PxPy_nonneg ac) (auto simp: space_pair_measure) |
49803 | 564 |
|
565 |
have X: "X = fst \<circ> (\<lambda>x. (X x, Y x))" and Y: "Y = snd \<circ> (\<lambda>x. (X x, Y x))" |
|
566 |
by auto |
|
567 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
568 |
have "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)))" |
49803 | 569 |
using finite_entropy_integrable[OF Fxy] |
570 |
using finite_entropy_integrable_transform[OF Fx Pxy, of fst] |
|
571 |
using finite_entropy_integrable_transform[OF Fy Pxy, of snd] |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
572 |
by (simp add: Pxy_nn) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
573 |
moreover have "f \<in> borel_measurable (S \<Otimes>\<^sub>M T)" |
49803 | 574 |
unfolding f_def using Px Py Pxy |
575 |
by (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'' |
|
576 |
intro!: borel_measurable_times borel_measurable_log borel_measurable_divide) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
577 |
ultimately have int: "integrable (S \<Otimes>\<^sub>M T) f" |
49803 | 578 |
apply (rule integrable_cong_AE_imp) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
579 |
using A B |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
580 |
by (auto simp: f_def log_divide_eq log_mult_eq field_simps space_pair_measure Px_nn Py_nn Pxy_nn |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
581 |
less_le) |
49803 | 582 |
|
583 |
show "0 \<le> ?M" unfolding M |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
584 |
proof (intro ST.KL_density_density_nonneg) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
585 |
show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x))) " |
49803 | 586 |
unfolding distributed_distr_eq_density[OF Pxy, symmetric] |
587 |
using distributed_measurable[OF Pxy] by (rule prob_space_distr) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
588 |
show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x))))" |
49803 | 589 |
unfolding distr_eq[symmetric] by unfold_locales |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
590 |
show "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
591 |
using int unfolding f_def . |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
592 |
qed (insert ac, auto simp: b_gt_1 Px_nn Py_nn Pxy_nn space_pair_measure) |
49803 | 593 |
qed |
594 |
||
595 |
lemma (in information_space) |
|
596 |
fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" |
|
47694 | 597 |
assumes "sigma_finite_measure S" "sigma_finite_measure T" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
598 |
assumes Px: "distributed M S X Px" and Px_nn: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
599 |
and Py: "distributed M T Y Py" and Py_nn: "\<And>y. y \<in> space T \<Longrightarrow> 0 \<le> Py y" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
600 |
and Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
601 |
and Pxy_nn: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)" |
47694 | 602 |
defines "f \<equiv> \<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
603 |
shows mutual_information_distr: "mutual_information b S T X Y = integral\<^sup>L (S \<Otimes>\<^sub>M T) f" (is "?M = ?R") |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
604 |
and mutual_information_nonneg: "integrable (S \<Otimes>\<^sub>M T) f \<Longrightarrow> 0 \<le> mutual_information b S T X Y" |
40859 | 605 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
606 |
have X[measurable]: "random_variable S X" |
47694 | 607 |
using Px by (auto simp: distributed_def) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
608 |
have Y[measurable]: "random_variable T Y" |
47694 | 609 |
using Py by (auto simp: distributed_def) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
610 |
have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
611 |
using Px Px_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
612 |
have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
613 |
using Py Py_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
614 |
have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
615 |
using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
616 |
|
47694 | 617 |
interpret S: sigma_finite_measure S by fact |
618 |
interpret T: sigma_finite_measure T by fact |
|
619 |
interpret ST: pair_sigma_finite S T .. |
|
620 |
interpret X: prob_space "distr M S X" using X by (rule prob_space_distr) |
|
621 |
interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr) |
|
622 |
interpret XY: pair_prob_space "distr M S X" "distr M T Y" .. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
623 |
let ?P = "S \<Otimes>\<^sub>M T" |
47694 | 624 |
let ?D = "distr M ?P (\<lambda>x. (X x, Y x))" |
625 |
||
626 |
{ fix A assume "A \<in> sets S" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
627 |
with X[THEN measurable_space] Y[THEN measurable_space] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
628 |
have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
629 |
by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } |
47694 | 630 |
note marginal_eq1 = this |
631 |
{ fix A assume "A \<in> sets T" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
632 |
with X[THEN measurable_space] Y[THEN measurable_space] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
633 |
have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
634 |
by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } |
47694 | 635 |
note marginal_eq2 = this |
636 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
637 |
have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
638 |
unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] |
47694 | 639 |
proof (subst pair_measure_density) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
640 |
show "(\<lambda>x. ennreal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ennreal (Py y)) \<in> borel_measurable T" |
47694 | 641 |
using Px Py by (auto simp: distributed_def) |
642 |
show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] .. |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
643 |
show "density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). ennreal (Px x) * ennreal (Py y)) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
644 |
density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
645 |
using Px_nn Py_nn by (auto intro!: density_cong simp: distributed_def ennreal_mult space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
646 |
qed fact |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
647 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
648 |
have M: "?M = KL_divergence b (density ?P (\<lambda>x. ennreal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ennreal (Pxy x)))" |
47694 | 649 |
unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] .. |
650 |
||
651 |
from Px Py have f: "(\<lambda>x. Px (fst x) * Py (snd x)) \<in> borel_measurable ?P" |
|
652 |
by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'') |
|
653 |
have PxPy_nonneg: "AE x in ?P. 0 \<le> Px (fst x) * Py (snd x)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
654 |
using Px_nn Py_nn by (auto simp: space_pair_measure) |
47694 | 655 |
|
656 |
have "(AE x in ?P. Px (fst x) = 0 \<longrightarrow> Pxy x = 0)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
657 |
by (rule subdensity_real[OF measurable_fst Pxy Px]) (insert Px_nn Pxy_nn, auto simp: space_pair_measure) |
47694 | 658 |
moreover |
659 |
have "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
660 |
by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure) |
47694 | 661 |
ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
662 |
by auto |
47694 | 663 |
|
664 |
show "?M = ?R" |
|
665 |
unfolding M f_def |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
666 |
using b_gt_1 f PxPy_nonneg ac Pxy_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
667 |
by (intro ST.KL_density_density) (auto simp: space_pair_measure) |
47694 | 668 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
669 |
assume int: "integrable (S \<Otimes>\<^sub>M T) f" |
47694 | 670 |
show "0 \<le> ?M" unfolding M |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
671 |
proof (intro ST.KL_density_density_nonneg) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
672 |
show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x))) " |
47694 | 673 |
unfolding distributed_distr_eq_density[OF Pxy, symmetric] |
674 |
using distributed_measurable[OF Pxy] by (rule prob_space_distr) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
675 |
show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Px (fst x) * Py (snd x))))" |
47694 | 676 |
unfolding distr_eq[symmetric] by unfold_locales |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
677 |
show "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
678 |
using int unfolding f_def . |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
679 |
qed (insert ac, auto simp: b_gt_1 Px_nn Py_nn Pxy_nn space_pair_measure) |
40859 | 680 |
qed |
681 |
||
682 |
lemma (in information_space) |
|
47694 | 683 |
fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" |
684 |
assumes "sigma_finite_measure S" "sigma_finite_measure T" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
685 |
assumes Px[measurable]: "distributed M S X Px" and Px_nn: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
686 |
and Py[measurable]: "distributed M T Y Py" and Py_nn: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
687 |
and Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
688 |
and Pxy_nn: "\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x" |
47694 | 689 |
assumes ae: "AE x in S. AE y in T. Pxy (x, y) = Px x * Py y" |
690 |
shows mutual_information_eq_0: "mutual_information b S T X Y = 0" |
|
36624 | 691 |
proof - |
47694 | 692 |
interpret S: sigma_finite_measure S by fact |
693 |
interpret T: sigma_finite_measure T by fact |
|
694 |
interpret ST: pair_sigma_finite S T .. |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
695 |
note |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
696 |
distributed_real_measurable[OF Px_nn Px, measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
697 |
distributed_real_measurable[OF Py_nn Py, measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
698 |
distributed_real_measurable[OF Pxy_nn Pxy, measurable] |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
699 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
700 |
have "AE x in S \<Otimes>\<^sub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
701 |
by (rule subdensity_real[OF measurable_fst Pxy Px]) (auto simp: Px_nn Pxy_nn space_pair_measure) |
47694 | 702 |
moreover |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
703 |
have "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
704 |
by (rule subdensity_real[OF measurable_snd Pxy Py]) (auto simp: Py_nn Pxy_nn space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
705 |
moreover |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
706 |
have "AE x in S \<Otimes>\<^sub>M T. Pxy x = Px (fst x) * Py (snd x)" |
47694 | 707 |
by (intro ST.AE_pair_measure) (auto simp: ae intro!: measurable_snd'' measurable_fst'') |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
708 |
ultimately have "AE x in S \<Otimes>\<^sub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0" |
47694 | 709 |
by eventually_elim simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
710 |
then have "(\<integral>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \<partial>(S \<Otimes>\<^sub>M T)) = (\<integral>x. 0 \<partial>(S \<Otimes>\<^sub>M T))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
711 |
by (intro integral_cong_AE) auto |
47694 | 712 |
then show ?thesis |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
713 |
by (subst mutual_information_distr[OF assms(1-8)]) (auto simp add: space_pair_measure) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
714 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
715 |
|
47694 | 716 |
lemma (in information_space) mutual_information_simple_distributed: |
717 |
assumes X: "simple_distributed M X Px" and Y: "simple_distributed M Y Py" |
|
718 |
assumes XY: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy" |
|
719 |
shows "\<I>(X ; Y) = (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x))`space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
720 |
proof (subst mutual_information_distr[OF _ _ simple_distributed[OF X] _ simple_distributed[OF Y] _ simple_distributed_joint[OF XY]]) |
47694 | 721 |
note fin = simple_distributed_joint_finite[OF XY, simp] |
722 |
show "sigma_finite_measure (count_space (X ` space M))" |
|
723 |
by (simp add: sigma_finite_measure_count_space_finite) |
|
724 |
show "sigma_finite_measure (count_space (Y ` space M))" |
|
725 |
by (simp add: sigma_finite_measure_count_space_finite) |
|
726 |
let ?Pxy = "\<lambda>x. (if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x else 0)" |
|
727 |
let ?f = "\<lambda>x. ?Pxy x * log b (?Pxy x / (Px (fst x) * Py (snd x)))" |
|
728 |
have "\<And>x. ?f x = (if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) else 0)" |
|
729 |
by auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
730 |
with fin show "(\<integral> x. ?f x \<partial>(count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M))) = |
47694 | 731 |
(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))" |
64267 | 732 |
by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite sum.If_cases split_beta' |
733 |
intro!: sum.cong) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
734 |
qed (insert X Y XY, auto simp: simple_distributed_def) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
735 |
|
47694 | 736 |
lemma (in information_space) |
737 |
fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" |
|
738 |
assumes Px: "simple_distributed M X Px" and Py: "simple_distributed M Y Py" |
|
739 |
assumes Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy" |
|
740 |
assumes ae: "\<forall>x\<in>space M. Pxy (X x, Y x) = Px (X x) * Py (Y x)" |
|
741 |
shows mutual_information_eq_0_simple: "\<I>(X ; Y) = 0" |
|
742 |
proof (subst mutual_information_simple_distributed[OF Px Py Pxy]) |
|
743 |
have "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = |
|
744 |
(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. 0)" |
|
64267 | 745 |
by (intro sum.cong) (auto simp: ae) |
47694 | 746 |
then show "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. |
747 |
Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = 0" by simp |
|
748 |
qed |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
749 |
|
61808 | 750 |
subsection \<open>Entropy\<close> |
39097 | 751 |
|
47694 | 752 |
definition (in prob_space) entropy :: "real \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> real" where |
753 |
"entropy b S X = - KL_divergence b S (distr M S X)" |
|
754 |
||
40859 | 755 |
abbreviation (in information_space) |
756 |
entropy_Pow ("\<H>'(_')") where |
|
47694 | 757 |
"\<H>(X) \<equiv> entropy b (count_space (X`space M)) X" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41833
diff
changeset
|
758 |
|
49791 | 759 |
lemma (in prob_space) distributed_RN_deriv: |
760 |
assumes X: "distributed M S X Px" |
|
761 |
shows "AE x in S. RN_deriv S (density S Px) x = Px x" |
|
762 |
proof - |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
763 |
have "distributed M S X (RN_deriv S (density S Px))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
764 |
by (metis RN_derivI assms borel_measurable_RN_deriv distributed_def) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
765 |
then show ?thesis |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
766 |
using assms distributed_unique by blast |
49791 | 767 |
qed |
768 |
||
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
769 |
lemma (in information_space) |
47694 | 770 |
fixes X :: "'a \<Rightarrow> 'b" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
771 |
assumes X[measurable]: "distributed M MX X f" and nn: "\<And>x. x \<in> space MX \<Longrightarrow> 0 \<le> f x" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
772 |
shows entropy_distr: "entropy b MX X = - (\<integral>x. f x * log b (f x) \<partial>MX)" (is ?eq) |
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
773 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
774 |
note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] |
49791 | 775 |
note ae = distributed_RN_deriv[OF X] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
776 |
note distributed_real_measurable[OF nn X, measurable] |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
777 |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
778 |
have ae_eq: "AE x in distr M MX X. log b (enn2real (RN_deriv MX (distr M MX X) x)) = log b (f x)" |
49785 | 779 |
unfolding distributed_distr_eq_density[OF X] |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
780 |
using D ae by (auto simp: AE_density) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
781 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
782 |
have int_eq: "(\<integral> x. f x * log b (f x) \<partial>MX) = (\<integral> x. log b (f x) \<partial>distr M MX X)" |
49785 | 783 |
unfolding distributed_distr_eq_density[OF X] |
784 |
using D |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
785 |
by (subst integral_density) (auto simp: nn) |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
786 |
|
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
787 |
show ?eq |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
788 |
unfolding entropy_def KL_divergence_def entropy_density_def comp_def int_eq neg_equal_iff_equal |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
789 |
using ae_eq by (intro integral_cong_AE) (auto simp: nn) |
49786 | 790 |
qed |
791 |
||
792 |
lemma (in information_space) entropy_le: |
|
793 |
fixes Px :: "'b \<Rightarrow> real" and MX :: "'b measure" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
794 |
assumes X[measurable]: "distributed M MX X Px" and Px_nn[simp]: "\<And>x. x \<in> space MX \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
795 |
and fin: "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> top" |
49786 | 796 |
and int: "integrable MX (\<lambda>x. - Px x * log b (Px x))" |
797 |
shows "entropy b MX X \<le> log b (measure MX {x \<in> space MX. Px x \<noteq> 0})" |
|
798 |
proof - |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
799 |
note Px = distributed_borel_measurable[OF X] |
49786 | 800 |
interpret X: prob_space "distr M MX X" |
801 |
using distributed_measurable[OF X] by (rule prob_space_distr) |
|
802 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
803 |
have " - log b (measure MX {x \<in> space MX. Px x \<noteq> 0}) = |
49786 | 804 |
- log b (\<integral> x. indicator {x \<in> space MX. Px x \<noteq> 0} x \<partial>MX)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
805 |
using Px Px_nn fin by (auto simp: measure_def) |
49786 | 806 |
also have "- log b (\<integral> x. indicator {x \<in> space MX. Px x \<noteq> 0} x \<partial>MX) = - log b (\<integral> x. 1 / Px x \<partial>distr M MX X)" |
67982
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
807 |
proof - |
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
808 |
have "integral\<^sup>L MX (indicator {x \<in> space MX. Px x \<noteq> 0}) = LINT x|MX. Px x *\<^sub>R (1 / Px x)" |
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
809 |
by (rule Bochner_Integration.integral_cong) auto |
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
810 |
also have "... = LINT x|density MX (\<lambda>x. ennreal (Px x)). 1 / Px x" |
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
811 |
by (rule integral_density [symmetric]) (use Px Px_nn in auto) |
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
812 |
finally show ?thesis |
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
813 |
unfolding distributed_distr_eq_density[OF X] by simp |
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
814 |
qed |
49786 | 815 |
also have "\<dots> \<le> (\<integral> x. - log b (1 / Px x) \<partial>distr M MX X)" |
816 |
proof (rule X.jensens_inequality[of "\<lambda>x. 1 / Px x" "{0<..}" 0 1 "\<lambda>x. - log b x"]) |
|
817 |
show "AE x in distr M MX X. 1 / Px x \<in> {0<..}" |
|
818 |
unfolding distributed_distr_eq_density[OF X] |
|
819 |
using Px by (auto simp: AE_density) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
820 |
have [simp]: "\<And>x. x \<in> space MX \<Longrightarrow> ennreal (if Px x = 0 then 0 else 1) = indicator {x \<in> space MX. Px x \<noteq> 0} x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
821 |
by (auto simp: one_ennreal_def) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
822 |
have "(\<integral>\<^sup>+ x. ennreal (- (if Px x = 0 then 0 else 1)) \<partial>MX) = (\<integral>\<^sup>+ x. 0 \<partial>MX)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
823 |
by (intro nn_integral_cong) (auto simp: ennreal_neg) |
49786 | 824 |
then show "integrable (distr M MX X) (\<lambda>x. 1 / Px x)" |
825 |
unfolding distributed_distr_eq_density[OF X] using Px |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
826 |
by (auto simp: nn_integral_density real_integrable_def fin ennreal_neg ennreal_mult[symmetric] |
56996 | 827 |
cong: nn_integral_cong) |
49786 | 828 |
have "integrable MX (\<lambda>x. Px x * log b (1 / Px x)) = |
829 |
integrable MX (\<lambda>x. - Px x * log b (Px x))" |
|
830 |
using Px |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
831 |
by (intro integrable_cong_AE) (auto simp: log_divide_eq less_le) |
49786 | 832 |
then show "integrable (distr M MX X) (\<lambda>x. - log b (1 / Px x))" |
833 |
unfolding distributed_distr_eq_density[OF X] |
|
834 |
using Px int |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
835 |
by (subst integrable_real_density) auto |
49786 | 836 |
qed (auto simp: minus_log_convex[OF b_gt_1]) |
837 |
also have "\<dots> = (\<integral> x. log b (Px x) \<partial>distr M MX X)" |
|
838 |
unfolding distributed_distr_eq_density[OF X] using Px |
|
839 |
by (intro integral_cong_AE) (auto simp: AE_density log_divide_eq) |
|
840 |
also have "\<dots> = - entropy b MX X" |
|
841 |
unfolding distributed_distr_eq_density[OF X] using Px |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
842 |
by (subst entropy_distr[OF X]) (auto simp: integral_density) |
49786 | 843 |
finally show ?thesis |
844 |
by simp |
|
845 |
qed |
|
846 |
||
847 |
lemma (in information_space) entropy_le_space: |
|
848 |
fixes Px :: "'b \<Rightarrow> real" and MX :: "'b measure" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
849 |
assumes X: "distributed M MX X Px" and Px_nn[simp]: "\<And>x. x \<in> space MX \<Longrightarrow> 0 \<le> Px x" |
49786 | 850 |
and fin: "finite_measure MX" |
851 |
and int: "integrable MX (\<lambda>x. - Px x * log b (Px x))" |
|
852 |
shows "entropy b MX X \<le> log b (measure MX (space MX))" |
|
853 |
proof - |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
854 |
note Px = distributed_borel_measurable[OF X] |
49786 | 855 |
interpret finite_measure MX by fact |
856 |
have "entropy b MX X \<le> log b (measure MX {x \<in> space MX. Px x \<noteq> 0})" |
|
857 |
using int X by (intro entropy_le) auto |
|
858 |
also have "\<dots> \<le> log b (measure MX (space MX))" |
|
859 |
using Px distributed_imp_emeasure_nonzero[OF X] |
|
80034
95b4fb2b5359
New material and a bit of refactoring
paulson <lp15@cam.ac.uk>
parents:
79772
diff
changeset
|
860 |
by (intro Transcendental.log_mono) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
861 |
(auto intro!: finite_measure_mono b_gt_1 less_le[THEN iffD2] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
862 |
simp: emeasure_eq_measure cong: conj_cong) |
49786 | 863 |
finally show ?thesis . |
864 |
qed |
|
865 |
||
47694 | 866 |
lemma (in information_space) entropy_uniform: |
49785 | 867 |
assumes X: "distributed M MX X (\<lambda>x. indicator A x / measure MX A)" (is "distributed _ _ _ ?f") |
47694 | 868 |
shows "entropy b MX X = log b (measure MX A)" |
49785 | 869 |
proof (subst entropy_distr[OF X]) |
870 |
have [simp]: "emeasure MX A \<noteq> \<infinity>" |
|
871 |
using uniform_distributed_params[OF X] by (auto simp add: measure_def) |
|
872 |
have eq: "(\<integral> x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \<partial>MX) = |
|
873 |
(\<integral> x. (- log b (measure MX A) / measure MX A) * indicator A x \<partial>MX)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
874 |
using uniform_distributed_params[OF X] |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
875 |
by (intro Bochner_Integration.integral_cong) (auto split: split_indicator simp: log_divide_eq zero_less_measure_iff) |
49785 | 876 |
show "- (\<integral> x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \<partial>MX) = |
877 |
log b (measure MX A)" |
|
878 |
unfolding eq using uniform_distributed_params[OF X] |
|
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
879 |
by (subst Bochner_Integration.integral_mult_right) (auto simp: measure_def less_top[symmetric] intro!: integrable_real_indicator) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
880 |
qed simp |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
881 |
|
47694 | 882 |
lemma (in information_space) entropy_simple_distributed: |
49786 | 883 |
"simple_distributed M X f \<Longrightarrow> \<H>(X) = - (\<Sum>x\<in>X`space M. f x * log b (f x))" |
884 |
by (subst entropy_distr[OF simple_distributed]) |
|
885 |
(auto simp add: lebesgue_integral_count_space_finite) |
|
39097 | 886 |
|
40859 | 887 |
lemma (in information_space) entropy_le_card_not_0: |
47694 | 888 |
assumes X: "simple_distributed M X f" |
889 |
shows "\<H>(X) \<le> log b (card (X ` space M \<inter> {x. f x \<noteq> 0}))" |
|
39097 | 890 |
proof - |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
891 |
let ?X = "count_space (X`space M)" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
892 |
have "\<H>(X) \<le> log b (measure ?X {x \<in> space ?X. f x \<noteq> 0})" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
893 |
by (rule entropy_le[OF simple_distributed[OF X]]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
894 |
(insert X, auto simp: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
895 |
also have "measure ?X {x \<in> space ?X. f x \<noteq> 0} = card (X ` space M \<inter> {x. f x \<noteq> 0})" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
896 |
by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def Int_def) |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
897 |
finally show ?thesis . |
39097 | 898 |
qed |
899 |
||
40859 | 900 |
lemma (in information_space) entropy_le_card: |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
901 |
assumes X: "simple_distributed M X f" |
40859 | 902 |
shows "\<H>(X) \<le> log b (real (card (X ` space M)))" |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
903 |
proof - |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
904 |
let ?X = "count_space (X`space M)" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
905 |
have "\<H>(X) \<le> log b (measure ?X (space ?X))" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
906 |
by (rule entropy_le_space[OF simple_distributed[OF X]]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
907 |
(insert X, auto simp: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space finite_measure_count_space) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
908 |
also have "measure ?X (space ?X) = card (X ` space M)" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
909 |
by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def) |
39097 | 910 |
finally show ?thesis . |
911 |
qed |
|
912 |
||
61808 | 913 |
subsection \<open>Conditional Mutual Information\<close> |
39097 | 914 |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
915 |
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:
41661
diff
changeset
|
916 |
"conditional_mutual_information b MX MY MZ X Y Z \<equiv> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
917 |
mutual_information b MX (MY \<Otimes>\<^sub>M MZ) X (\<lambda>x. (Y x, Z x)) - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
918 |
mutual_information b MX MZ X Z" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
919 |
|
40859 | 920 |
abbreviation (in information_space) |
921 |
conditional_mutual_information_Pow ("\<I>'( _ ; _ | _ ')") where |
|
36624 | 922 |
"\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b |
47694 | 923 |
(count_space (X ` space M)) (count_space (Y ` space M)) (count_space (Z ` space M)) X Y Z" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
924 |
|
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
925 |
lemma (in information_space) |
47694 | 926 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P" |
50003 | 927 |
assumes Px[measurable]: "distributed M S X Px" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
928 |
and Px_nn[simp]: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x" |
50003 | 929 |
assumes Pz[measurable]: "distributed M P Z Pz" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
930 |
and Pz_nn[simp]: "\<And>z. z \<in> space P \<Longrightarrow> 0 \<le> Pz z" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
931 |
assumes Pyz[measurable]: "distributed M (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x)) Pyz" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
932 |
and Pyz_nn[simp]: "\<And>y z. y \<in> space T \<Longrightarrow> z \<in> space P \<Longrightarrow> 0 \<le> Pyz (y, z)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
933 |
assumes Pxz[measurable]: "distributed M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) Pxz" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
934 |
and Pxz_nn[simp]: "\<And>x z. x \<in> space S \<Longrightarrow> z \<in> space P \<Longrightarrow> 0 \<le> Pxz (x, z)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
935 |
assumes Pxyz[measurable]: "distributed M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
936 |
and Pxyz_nn[simp]: "\<And>x y z. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> z \<in> space P \<Longrightarrow> 0 \<le> Pxyz (x, y, z)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
937 |
assumes I1: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
938 |
assumes I2: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))" |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
939 |
shows conditional_mutual_information_generic_eq: "conditional_mutual_information b S T P X Y Z |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
940 |
= (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" (is "?eq") |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
941 |
and conditional_mutual_information_generic_nonneg: "0 \<le> conditional_mutual_information b S T P X Y Z" (is "?nonneg") |
40859 | 942 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
943 |
have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
944 |
using Px Px_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
945 |
have [measurable]: "Pz \<in> P \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
946 |
using Pz Pz_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
947 |
have measurable_Pyz[measurable]: "Pyz \<in> (T \<Otimes>\<^sub>M P) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
948 |
using Pyz Pyz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
949 |
have measurable_Pxz[measurable]: "Pxz \<in> (S \<Otimes>\<^sub>M P) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
950 |
using Pxz Pxz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
951 |
have measurable_Pxyz[measurable]: "Pxyz \<in> (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
952 |
using Pxyz Pxyz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
953 |
|
47694 | 954 |
interpret S: sigma_finite_measure S by fact |
955 |
interpret T: sigma_finite_measure T by fact |
|
956 |
interpret P: sigma_finite_measure P by fact |
|
957 |
interpret TP: pair_sigma_finite T P .. |
|
958 |
interpret SP: pair_sigma_finite S P .. |
|
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
959 |
interpret ST: pair_sigma_finite S T .. |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
960 |
interpret SPT: pair_sigma_finite "S \<Otimes>\<^sub>M P" T .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
961 |
interpret STP: pair_sigma_finite S "T \<Otimes>\<^sub>M P" .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
962 |
interpret TPS: pair_sigma_finite "T \<Otimes>\<^sub>M P" S .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
963 |
have TP: "sigma_finite_measure (T \<Otimes>\<^sub>M P)" .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
964 |
have SP: "sigma_finite_measure (S \<Otimes>\<^sub>M P)" .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
965 |
have YZ: "random_variable (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x))" |
47694 | 966 |
using Pyz by (simp add: distributed_measurable) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
967 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
968 |
from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
969 |
distr (distr M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). (x, z))" |
50003 | 970 |
by (simp add: comp_def distr_distr) |
40859 | 971 |
|
47694 | 972 |
have "mutual_information b S P X Z = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
973 |
(\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^sub>M P))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
974 |
by (rule mutual_information_distr[OF S P Px Px_nn Pz Pz_nn Pxz Pxz_nn]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
975 |
also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" |
47694 | 976 |
using b_gt_1 Pxz Px Pz |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
977 |
by (subst distributed_transform_integral[OF Pxyz _ Pxz _, where T="\<lambda>(x, y, z). (x, z)"]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
978 |
(auto simp: split_beta' space_pair_measure) |
47694 | 979 |
finally have mi_eq: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
980 |
"mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" . |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
981 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
982 |
have ae1: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
983 |
by (intro subdensity_real[of fst, OF _ Pxyz Px]) (auto simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
984 |
moreover have ae2: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
985 |
by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) (auto simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
986 |
moreover have ae3: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
987 |
by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) (auto simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
988 |
moreover have ae4: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
989 |
by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) (auto simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
990 |
ultimately have ae: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. |
47694 | 991 |
Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) - |
992 |
Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) = |
|
993 |
Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) " |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
994 |
using AE_space |
47694 | 995 |
proof eventually_elim |
60580 | 996 |
case (elim x) |
47694 | 997 |
show ?case |
40859 | 998 |
proof cases |
47694 | 999 |
assume "Pxyz x \<noteq> 0" |
60580 | 1000 |
with elim have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))" |
1001 |
"0 < Pyz (snd x)" "0 < Pxyz x" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1002 |
by (auto simp: space_pair_measure less_le) |
47694 | 1003 |
then show ?thesis |
56544 | 1004 |
using b_gt_1 by (simp add: log_simps less_imp_le field_simps) |
40859 | 1005 |
qed simp |
1006 |
qed |
|
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1007 |
with I1 I2 show ?eq |
40859 | 1008 |
unfolding conditional_mutual_information_def |
47694 | 1009 |
apply (subst mi_eq) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1010 |
apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz _ Pxyz]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1011 |
apply (auto simp: space_pair_measure) |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1012 |
apply (subst Bochner_Integration.integral_diff[symmetric]) |
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1013 |
apply (auto intro!: integral_cong_AE simp: split_beta' simp del: Bochner_Integration.integral_diff) |
47694 | 1014 |
done |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1015 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1016 |
let ?P = "density (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) Pxyz" |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1017 |
interpret P: prob_space ?P |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1018 |
unfolding distributed_distr_eq_density[OF Pxyz, symmetric] |
50003 | 1019 |
by (rule prob_space_distr) simp |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1020 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1021 |
let ?Q = "density (T \<Otimes>\<^sub>M P) Pyz" |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1022 |
interpret Q: prob_space ?Q |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1023 |
unfolding distributed_distr_eq_density[OF Pyz, symmetric] |
50003 | 1024 |
by (rule prob_space_distr) simp |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1025 |
|
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1026 |
let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1027 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1028 |
from subdensity_real[of snd, OF _ Pyz Pz _ AE_I2 AE_I2] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1029 |
have aeX1: "AE x in T \<Otimes>\<^sub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1030 |
by (auto simp: comp_def space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1031 |
have aeX2: "AE x in T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd x)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1032 |
using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1033 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1034 |
have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ennreal (Pxz (x, snd y)) \<partial>S) = ennreal (Pz (snd y))" |
49788
3c10763f5cb4
show and use distributed_swap and distributed_jointI
hoelzl
parents:
49787
diff
changeset
|
1035 |
using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1036 |
by (intro TP.AE_pair_measure) auto |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1037 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1038 |
have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1039 |
by (subst nn_integral_density) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1040 |
(auto intro!: nn_integral_mono simp: space_pair_measure ennreal_mult[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1041 |
also have "\<dots> = (\<integral>\<^sup>+(y, z). (\<integral>\<^sup>+ x. ennreal (Pxz (x, z)) * ennreal (Pyz (y, z) / Pz z) \<partial>S) \<partial>(T \<Otimes>\<^sub>M P))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1042 |
by (subst STP.nn_integral_snd[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1043 |
(auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1044 |
also have "\<dots> = (\<integral>\<^sup>+x. ennreal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1045 |
proof - |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1046 |
have D: "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1047 |
if "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "a \<in> space T \<and> b \<in> space P" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1048 |
"(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1049 |
for a b |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1050 |
using that by (subst nn_integral_multc) (auto split: prod.split simp: ennreal_mult[symmetric]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1051 |
show ?thesis |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1052 |
apply (rule nn_integral_cong_AE) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1053 |
using aeX1 aeX2 aeX3 |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1054 |
by (force simp add: space_pair_measure D) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1055 |
qed |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1056 |
also have "\<dots> = 1" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1057 |
using Q.emeasure_space_1 distributed_distr_eq_density[OF Pyz] |
56996 | 1058 |
by (subst nn_integral_density[symmetric]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1059 |
finally have le1: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> 1" . |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1060 |
also have "\<dots> < \<infinity>" by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1061 |
finally have fin: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1062 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1063 |
have pos: "(\<integral>\<^sup>+x. ?f x \<partial>?P) \<noteq> 0" |
56996 | 1064 |
apply (subst nn_integral_density) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1065 |
apply (simp_all add: split_beta') |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1066 |
proof |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1067 |
let ?g = "\<lambda>x. ennreal (Pxyz x) * (Pxz (fst x, snd (snd x)) * Pyz (snd x) / (Pz (snd (snd x)) * Pxyz x))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1068 |
assume "(\<integral>\<^sup>+x. ?g x \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)) = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1069 |
then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x = 0" |
56996 | 1070 |
by (intro nn_integral_0_iff_AE[THEN iffD1]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1071 |
then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1072 |
using ae1 ae2 ae3 ae4 |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1073 |
by (auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1074 |
then have "(\<integral>\<^sup>+ x. ennreal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0" |
56996 | 1075 |
by (subst nn_integral_cong_AE[of _ "\<lambda>x. 0"]) auto |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1076 |
with P.emeasure_space_1 show False |
56996 | 1077 |
by (subst (asm) emeasure_density) (auto cong: nn_integral_cong) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1078 |
qed |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1079 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1080 |
have neg: "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) = 0" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1081 |
by (subst nn_integral_0_iff_AE) (auto simp: space_pair_measure ennreal_neg) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1082 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1083 |
have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1084 |
apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ Bochner_Integration.integrable_diff[OF I1 I2]]) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1085 |
using ae |
50003 | 1086 |
apply (auto simp: split_beta') |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1087 |
done |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1088 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1089 |
have "- log b 1 \<le> - log b (integral\<^sup>L ?P ?f)" |
80034
95b4fb2b5359
New material and a bit of refactoring
paulson <lp15@cam.ac.uk>
parents:
79772
diff
changeset
|
1090 |
proof (intro le_imp_neg_le Transcendental.log_mono[OF b_gt_1]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1091 |
have If: "integrable ?P ?f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1092 |
unfolding real_integrable_def |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1093 |
proof (intro conjI) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1094 |
from neg show "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) \<noteq> \<infinity>" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1095 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1096 |
from fin show "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1097 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1098 |
qed simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1099 |
then have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) = (\<integral>x. ?f x \<partial>?P)" |
56996 | 1100 |
apply (rule nn_integral_eq_integral) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1101 |
apply (auto simp: space_pair_measure ennreal_neg) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1102 |
done |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1103 |
with pos le1 |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1104 |
show "0 < (\<integral>x. ?f x \<partial>?P)" "(\<integral>x. ?f x \<partial>?P) \<le> 1" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1105 |
by (simp_all add: one_ennreal.rep_eq zero_less_iff_neq_zero[symmetric]) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1106 |
qed |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1107 |
also have "- log b (integral\<^sup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)" |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1108 |
proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1109 |
show "AE x in ?P. ?f x \<in> {0<..}" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1110 |
unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]] |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1111 |
using ae1 ae2 ae3 ae4 |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1112 |
by (auto simp: space_pair_measure less_le) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1113 |
show "integrable ?P ?f" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1114 |
unfolding real_integrable_def |
50003 | 1115 |
using fin neg by (auto simp: split_beta') |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1116 |
have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. Pxyz x * - log b (?f x))" |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1117 |
apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3]) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1118 |
using ae1 ae2 ae3 ae4 |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1119 |
apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1120 |
less_le space_pair_measure) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1121 |
done |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1122 |
then |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1123 |
show "integrable ?P (\<lambda>x. - log b (?f x))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1124 |
by (subst integrable_real_density) (auto simp: space_pair_measure) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1125 |
qed (auto simp: b_gt_1 minus_log_convex) |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1126 |
also have "\<dots> = conditional_mutual_information b S T P X Y Z" |
61808 | 1127 |
unfolding \<open>?eq\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1128 |
apply (subst integral_real_density) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1129 |
apply simp |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1130 |
apply (force simp: space_pair_measure) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1131 |
apply simp |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1132 |
apply (intro integral_cong_AE) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1133 |
using ae1 ae2 ae3 ae4 |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1134 |
apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1135 |
space_pair_measure less_le) |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1136 |
done |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1137 |
finally show ?nonneg |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1138 |
by simp |
40859 | 1139 |
qed |
1140 |
||
49803 | 1141 |
lemma (in information_space) |
1142 |
fixes Px :: "_ \<Rightarrow> real" |
|
1143 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P" |
|
1144 |
assumes Fx: "finite_entropy S X Px" |
|
1145 |
assumes Fz: "finite_entropy P Z Pz" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1146 |
assumes Fyz: "finite_entropy (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x)) Pyz" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1147 |
assumes Fxz: "finite_entropy (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) Pxz" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1148 |
assumes Fxyz: "finite_entropy (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz" |
49803 | 1149 |
shows conditional_mutual_information_generic_eq': "conditional_mutual_information b S T P X Y Z |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1150 |
= (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" (is "?eq") |
49803 | 1151 |
and conditional_mutual_information_generic_nonneg': "0 \<le> conditional_mutual_information b S T P X Y Z" (is "?nonneg") |
1152 |
proof - |
|
50003 | 1153 |
note Px = Fx[THEN finite_entropy_distributed, measurable] |
1154 |
note Pz = Fz[THEN finite_entropy_distributed, measurable] |
|
1155 |
note Pyz = Fyz[THEN finite_entropy_distributed, measurable] |
|
1156 |
note Pxz = Fxz[THEN finite_entropy_distributed, measurable] |
|
1157 |
note Pxyz = Fxyz[THEN finite_entropy_distributed, measurable] |
|
49803 | 1158 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1159 |
note Px_nn = Fx[THEN finite_entropy_nn] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1160 |
note Pz_nn = Fz[THEN finite_entropy_nn] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1161 |
note Pyz_nn = Fyz[THEN finite_entropy_nn] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1162 |
note Pxz_nn = Fxz[THEN finite_entropy_nn] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1163 |
note Pxyz_nn = Fxyz[THEN finite_entropy_nn] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1164 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1165 |
note Px' = Fx[THEN finite_entropy_measurable, measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1166 |
note Pz' = Fz[THEN finite_entropy_measurable, measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1167 |
note Pyz' = Fyz[THEN finite_entropy_measurable, measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1168 |
note Pxz' = Fxz[THEN finite_entropy_measurable, measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1169 |
note Pxyz' = Fxyz[THEN finite_entropy_measurable, measurable] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1170 |
|
49803 | 1171 |
interpret S: sigma_finite_measure S by fact |
1172 |
interpret T: sigma_finite_measure T by fact |
|
1173 |
interpret P: sigma_finite_measure P by fact |
|
1174 |
interpret TP: pair_sigma_finite T P .. |
|
1175 |
interpret SP: pair_sigma_finite S P .. |
|
1176 |
interpret ST: pair_sigma_finite S T .. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1177 |
interpret SPT: pair_sigma_finite "S \<Otimes>\<^sub>M P" T .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1178 |
interpret STP: pair_sigma_finite S "T \<Otimes>\<^sub>M P" .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1179 |
interpret TPS: pair_sigma_finite "T \<Otimes>\<^sub>M P" S .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1180 |
have TP: "sigma_finite_measure (T \<Otimes>\<^sub>M P)" .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1181 |
have SP: "sigma_finite_measure (S \<Otimes>\<^sub>M P)" .. |
49803 | 1182 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1183 |
from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1184 |
distr (distr M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). (x, z))" |
50003 | 1185 |
by (simp add: distr_distr comp_def) |
49803 | 1186 |
|
1187 |
have "mutual_information b S P X Z = |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1188 |
(\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^sub>M P))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1189 |
using Px Px_nn Pz Pz_nn Pxz Pxz_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1190 |
by (rule mutual_information_distr[OF S P]) (auto simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1191 |
also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1192 |
using b_gt_1 Pxz Pxz_nn Pxyz Pxyz_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1193 |
by (subst distributed_transform_integral[OF Pxyz _ Pxz, where T="\<lambda>(x, y, z). (x, z)"]) |
50003 | 1194 |
(auto simp: split_beta') |
49803 | 1195 |
finally have mi_eq: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1196 |
"mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" . |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1197 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1198 |
have ae1: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1199 |
by (intro subdensity_finite_entropy[of fst, OF _ Fxyz Fx]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1200 |
moreover have ae2: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1201 |
by (intro subdensity_finite_entropy[of "\<lambda>x. snd (snd x)", OF _ Fxyz Fz]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1202 |
moreover have ae3: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1203 |
by (intro subdensity_finite_entropy[of "\<lambda>x. (fst x, snd (snd x))", OF _ Fxyz Fxz]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1204 |
moreover have ae4: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1205 |
by (intro subdensity_finite_entropy[of snd, OF _ Fxyz Fyz]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1206 |
ultimately have ae: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. |
49803 | 1207 |
Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) - |
1208 |
Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) = |
|
1209 |
Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) " |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1210 |
using AE_space |
49803 | 1211 |
proof eventually_elim |
60580 | 1212 |
case (elim x) |
49803 | 1213 |
show ?case |
1214 |
proof cases |
|
1215 |
assume "Pxyz x \<noteq> 0" |
|
60580 | 1216 |
with elim have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))" |
1217 |
"0 < Pyz (snd x)" "0 < Pxyz x" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1218 |
using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1219 |
by (auto simp: space_pair_measure less_le) |
49803 | 1220 |
then show ?thesis |
56544 | 1221 |
using b_gt_1 by (simp add: log_simps less_imp_le field_simps) |
49803 | 1222 |
qed simp |
1223 |
qed |
|
1224 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1225 |
have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) |
49803 | 1226 |
(\<lambda>x. Pxyz x * log b (Pxyz x) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pyz (snd x)))" |
1227 |
using finite_entropy_integrable[OF Fxyz] |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1228 |
using finite_entropy_integrable_transform[OF Fx Pxyz Pxyz_nn, of fst] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1229 |
using finite_entropy_integrable_transform[OF Fyz Pxyz Pxyz_nn, of snd] |
49803 | 1230 |
by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1231 |
moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z)))) \<in> borel_measurable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)" |
50003 | 1232 |
using Pxyz Px Pyz by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1233 |
ultimately have I1: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))" |
49803 | 1234 |
apply (rule integrable_cong_AE_imp) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1235 |
using ae1 ae4 Px_nn Pyz_nn Pxyz_nn |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1236 |
by (auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff space_pair_measure less_le) |
49803 | 1237 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1238 |
have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) |
49803 | 1239 |
(\<lambda>x. Pxyz x * log b (Pxz (fst x, snd (snd x))) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pz (snd (snd x))))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1240 |
using finite_entropy_integrable_transform[OF Fxz Pxyz Pxyz_nn, of "\<lambda>x. (fst x, snd (snd x))"] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1241 |
using finite_entropy_integrable_transform[OF Fx Pxyz Pxyz_nn, of fst] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1242 |
using finite_entropy_integrable_transform[OF Fz Pxyz Pxyz_nn, of "snd \<circ> snd"] |
50003 | 1243 |
by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1244 |
moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z))) \<in> borel_measurable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)" |
49803 | 1245 |
using Pxyz Px Pz |
50003 | 1246 |
by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1247 |
ultimately have I2: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))" |
49803 | 1248 |
apply (rule integrable_cong_AE_imp) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1249 |
using ae1 ae2 ae3 ae4 Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1250 |
by (auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff less_le space_pair_measure) |
49803 | 1251 |
|
1252 |
from ae I1 I2 show ?eq |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1253 |
unfolding conditional_mutual_information_def mi_eq |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1254 |
apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz Pyz_nn Pxyz Pxyz_nn]; simp add: space_pair_measure) |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1255 |
apply (subst Bochner_Integration.integral_diff[symmetric]) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1256 |
apply (auto intro!: integral_cong_AE simp: split_beta' simp del: Bochner_Integration.integral_diff) |
49803 | 1257 |
done |
1258 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1259 |
let ?P = "density (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) Pxyz" |
49803 | 1260 |
interpret P: prob_space ?P |
50003 | 1261 |
unfolding distributed_distr_eq_density[OF Pxyz, symmetric] by (rule prob_space_distr) simp |
49803 | 1262 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1263 |
let ?Q = "density (T \<Otimes>\<^sub>M P) Pyz" |
49803 | 1264 |
interpret Q: prob_space ?Q |
50003 | 1265 |
unfolding distributed_distr_eq_density[OF Pyz, symmetric] by (rule prob_space_distr) simp |
49803 | 1266 |
|
1267 |
let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)" |
|
1268 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1269 |
from subdensity_finite_entropy[of snd, OF _ Fyz Fz] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1270 |
have aeX1: "AE x in T \<Otimes>\<^sub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1271 |
have aeX2: "AE x in T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd x)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1272 |
using Pz by (intro TP.AE_pair_measure) (auto intro: Pz_nn) |
49803 | 1273 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1274 |
have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ennreal (Pxz (x, snd y)) \<partial>S) = ennreal (Pz (snd y))" |
49803 | 1275 |
using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1276 |
by (intro TP.AE_pair_measure) (auto ) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1277 |
have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1278 |
using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1279 |
by (subst nn_integral_density) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1280 |
(auto intro!: nn_integral_mono simp: space_pair_measure ennreal_mult[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1281 |
also have "\<dots> = (\<integral>\<^sup>+(y, z). \<integral>\<^sup>+ x. ennreal (Pxz (x, z)) * ennreal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^sub>M P)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1282 |
using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1283 |
by (subst STP.nn_integral_snd[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1284 |
(auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1285 |
also have "\<dots> = (\<integral>\<^sup>+x. ennreal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1286 |
proof - |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1287 |
have *: "(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \<partial>S) = ennreal (Pyz (a, b))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1288 |
if "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1289 |
"(\<integral>\<^sup>+ x. ennreal (Pxz (x, b)) \<partial>S) = ennreal (Pz b)" for a b |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1290 |
using Pyz_nn[of "(a,b)"] that |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1291 |
by (subst nn_integral_multc) (auto simp: space_pair_measure ennreal_mult[symmetric]) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1292 |
show ?thesis |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1293 |
using aeX1 aeX2 aeX3 AE_space |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1294 |
by (force simp: * space_pair_measure intro: nn_integral_cong_AE) |
49803 | 1295 |
qed |
1296 |
also have "\<dots> = 1" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1297 |
using Q.emeasure_space_1 Pyz_nn distributed_distr_eq_density[OF Pyz] |
56996 | 1298 |
by (subst nn_integral_density[symmetric]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1299 |
finally have le1: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> 1" . |
49803 | 1300 |
also have "\<dots> < \<infinity>" by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1301 |
finally have fin: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp |
49803 | 1302 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1303 |
have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1304 |
using Pxyz_nn |
56996 | 1305 |
apply (subst nn_integral_density) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1306 |
apply (simp_all add: split_beta' ennreal_mult'[symmetric] cong: nn_integral_cong) |
49803 | 1307 |
proof |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1308 |
let ?g = "\<lambda>x. ennreal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1309 |
assume "(\<integral>\<^sup>+ x. ?g x \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)) = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1310 |
then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1311 |
by (intro nn_integral_0_iff_AE[THEN iffD1]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1312 |
then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1313 |
using ae1 ae2 ae3 ae4 |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1314 |
by (insert Px_nn Pz_nn Pxz_nn Pyz_nn, |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1315 |
auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1316 |
then have "(\<integral>\<^sup>+ x. ennreal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0" |
56996 | 1317 |
by (subst nn_integral_cong_AE[of _ "\<lambda>x. 0"]) auto |
49803 | 1318 |
with P.emeasure_space_1 show False |
56996 | 1319 |
by (subst (asm) emeasure_density) (auto cong: nn_integral_cong) |
49803 | 1320 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1321 |
then have pos: "0 < (\<integral>\<^sup>+ x. ?f x \<partial>?P)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1322 |
by (simp add: zero_less_iff_neq_zero) |
49803 | 1323 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1324 |
have neg: "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1325 |
using Pz_nn Pxz_nn Pyz_nn Pxyz_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1326 |
by (intro nn_integral_0_iff_AE[THEN iffD2]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1327 |
(auto simp: split_beta' AE_density space_pair_measure intro!: AE_I2 ennreal_neg) |
49803 | 1328 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1329 |
have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1330 |
apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ Bochner_Integration.integrable_diff[OF I1 I2]]) |
49803 | 1331 |
using ae |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1332 |
by (auto simp: split_beta') |
49803 | 1333 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1334 |
have "- log b 1 \<le> - log b (integral\<^sup>L ?P ?f)" |
80034
95b4fb2b5359
New material and a bit of refactoring
paulson <lp15@cam.ac.uk>
parents:
79772
diff
changeset
|
1335 |
proof (intro le_imp_neg_le Transcendental.log_mono[OF b_gt_1]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1336 |
have If: "integrable ?P ?f" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1337 |
using neg fin by (force simp add: real_integrable_def) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1338 |
then have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) = (\<integral>x. ?f x \<partial>?P)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1339 |
using Pz_nn Pxz_nn Pyz_nn Pxyz_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1340 |
by (intro nn_integral_eq_integral) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1341 |
(auto simp: AE_density space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1342 |
with pos le1 |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1343 |
show "0 < (\<integral>x. ?f x \<partial>?P)" "(\<integral>x. ?f x \<partial>?P) \<le> 1" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1344 |
by (simp_all add: ) |
49803 | 1345 |
qed |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1346 |
also have "- log b (integral\<^sup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)" |
49803 | 1347 |
proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) |
1348 |
show "AE x in ?P. ?f x \<in> {0<..}" |
|
1349 |
unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]] |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1350 |
using ae1 ae2 ae3 ae4 |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1351 |
by (insert Pxyz_nn Pyz_nn Pz_nn Pxz_nn, auto simp: space_pair_measure less_le) |
49803 | 1352 |
show "integrable ?P ?f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1353 |
unfolding real_integrable_def |
50003 | 1354 |
using fin neg by (auto simp: split_beta') |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1355 |
have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. Pxyz x * - log b (?f x))" |
49803 | 1356 |
apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3]) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1357 |
using Pz_nn Pxz_nn Pyz_nn Pxyz_nn ae2 ae3 ae4 |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1358 |
by (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1359 |
zero_less_divide_iff field_simps space_pair_measure less_le) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1360 |
then |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1361 |
show "integrable ?P (\<lambda>x. - log b (?f x))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1362 |
using Pxyz_nn by (auto simp: integrable_real_density) |
49803 | 1363 |
qed (auto simp: b_gt_1 minus_log_convex) |
1364 |
also have "\<dots> = conditional_mutual_information b S T P X Y Z" |
|
61808 | 1365 |
unfolding \<open>?eq\<close> |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1366 |
using Pz_nn Pxz_nn Pyz_nn Pxyz_nn |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1367 |
apply (subst integral_real_density) |
50003 | 1368 |
apply simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1369 |
apply simp |
50003 | 1370 |
apply simp |
49803 | 1371 |
apply (intro integral_cong_AE) |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1372 |
using ae1 ae2 ae3 ae4 |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1373 |
apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1374 |
field_simps space_pair_measure less_le integral_cong_AE) |
49803 | 1375 |
done |
1376 |
finally show ?nonneg |
|
1377 |
by simp |
|
1378 |
qed |
|
1379 |
||
40859 | 1380 |
lemma (in information_space) conditional_mutual_information_eq: |
47694 | 1381 |
assumes Pz: "simple_distributed M Z Pz" |
1382 |
assumes Pyz: "simple_distributed M (\<lambda>x. (Y x, Z x)) Pyz" |
|
1383 |
assumes Pxz: "simple_distributed M (\<lambda>x. (X x, Z x)) Pxz" |
|
1384 |
assumes Pxyz: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Pxyz" |
|
1385 |
shows "\<I>(X ; Y | Z) = |
|
1386 |
(\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x))`space M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1387 |
proof (subst conditional_mutual_information_generic_eq[OF _ _ _ _ _ |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1388 |
simple_distributed[OF Pz] _ simple_distributed_joint[OF Pyz] _ simple_distributed_joint[OF Pxz] _ |
47694 | 1389 |
simple_distributed_joint2[OF Pxyz]]) |
1390 |
note simple_distributed_joint2_finite[OF Pxyz, simp] |
|
1391 |
show "sigma_finite_measure (count_space (X ` space M))" |
|
1392 |
by (simp add: sigma_finite_measure_count_space_finite) |
|
1393 |
show "sigma_finite_measure (count_space (Y ` space M))" |
|
1394 |
by (simp add: sigma_finite_measure_count_space_finite) |
|
1395 |
show "sigma_finite_measure (count_space (Z ` space M))" |
|
1396 |
by (simp add: sigma_finite_measure_count_space_finite) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1397 |
have "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) \<Otimes>\<^sub>M count_space (Z ` space M) = |
47694 | 1398 |
count_space (X`space M \<times> Y`space M \<times> Z`space M)" |
1399 |
(is "?P = ?C") |
|
1400 |
by (simp add: pair_measure_count_space) |
|
40859 | 1401 |
|
47694 | 1402 |
let ?Px = "\<lambda>x. measure M (X -` {x} \<inter> space M)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1403 |
have "(\<lambda>x. (X x, Z x)) \<in> measurable M (count_space (X ` space M) \<Otimes>\<^sub>M count_space (Z ` space M))" |
47694 | 1404 |
using simple_distributed_joint[OF Pxz] by (rule distributed_measurable) |
1405 |
from measurable_comp[OF this measurable_fst] |
|
1406 |
have "random_variable (count_space (X ` space M)) X" |
|
1407 |
by (simp add: comp_def) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1408 |
then have "simple_function M X" |
50002
ce0d316b5b44
add measurability prover; add support for Borel sets
hoelzl
parents:
49999
diff
changeset
|
1409 |
unfolding simple_function_def by (auto simp: measurable_count_space_eq2) |
47694 | 1410 |
then have "simple_distributed M X ?Px" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1411 |
by (rule simple_distributedI) (auto simp: measure_nonneg) |
47694 | 1412 |
then show "distributed M (count_space (X ` space M)) X ?Px" |
1413 |
by (rule simple_distributed) |
|
1414 |
||
1415 |
let ?f = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M then Pxyz x else 0)" |
|
1416 |
let ?g = "(\<lambda>x. if x \<in> (\<lambda>x. (Y x, Z x)) ` space M then Pyz x else 0)" |
|
1417 |
let ?h = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Z x)) ` space M then Pxz x else 0)" |
|
1418 |
show |
|
1419 |
"integrable ?P (\<lambda>(x, y, z). ?f (x, y, z) * log b (?f (x, y, z) / (?Px x * ?g (y, z))))" |
|
1420 |
"integrable ?P (\<lambda>(x, y, z). ?f (x, y, z) * log b (?h (x, z) / (?Px x * Pz z)))" |
|
1421 |
by (auto intro!: integrable_count_space simp: pair_measure_count_space) |
|
1422 |
let ?i = "\<lambda>x y z. ?f (x, y, z) * log b (?f (x, y, z) / (?h (x, z) * (?g (y, z) / Pz z)))" |
|
1423 |
let ?j = "\<lambda>x y z. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z)))" |
|
1424 |
have "(\<lambda>(x, y, z). ?i x y z) = (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M then ?j (fst x) (fst (snd x)) (snd (snd x)) else 0)" |
|
1425 |
by (auto intro!: ext) |
|
1426 |
then show "(\<integral> (x, y, z). ?i x y z \<partial>?P) = (\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x)) ` space M. ?j x y z)" |
|
64267 | 1427 |
by (auto intro!: sum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite sum.If_cases split_beta') |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1428 |
qed (insert Pz Pyz Pxz Pxyz, auto intro: measure_nonneg) |
36624 | 1429 |
|
47694 | 1430 |
lemma (in information_space) conditional_mutual_information_nonneg: |
1431 |
assumes X: "simple_function M X" and Y: "simple_function M Y" and Z: "simple_function M Z" |
|
1432 |
shows "0 \<le> \<I>(X ; Y | Z)" |
|
1433 |
proof - |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1434 |
have [simp]: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) \<Otimes>\<^sub>M count_space (Z ` space M) = |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1435 |
count_space (X`space M \<times> Y`space M \<times> Z`space M)" |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1436 |
by (simp add: pair_measure_count_space X Y Z simple_functionD) |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1437 |
note sf = sigma_finite_measure_count_space_finite[OF simple_functionD(1)] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1438 |
note sd = simple_distributedI[OF _ _ refl] |
49787
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1439 |
note sp = simple_function_Pair |
d8de705b48d4
rule to show that conditional mutual information is non-negative in the continuous case
hoelzl
parents:
49786
diff
changeset
|
1440 |
show ?thesis |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1441 |
apply (rule conditional_mutual_information_generic_nonneg[OF sf[OF X] sf[OF Y] sf[OF Z]]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1442 |
apply (force intro: simple_distributed[OF sd[OF X]]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1443 |
apply simp |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1444 |
apply (force intro: simple_distributed[OF sd[OF Z]]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1445 |
apply simp |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1446 |
apply (force intro: simple_distributed_joint[OF sd[OF sp[OF Y Z]]]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1447 |
apply simp |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1448 |
apply (force intro: simple_distributed_joint[OF sd[OF sp[OF X Z]]]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1449 |
apply simp |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1450 |
apply (fastforce intro: simple_distributed_joint2[OF sd[OF sp[OF X sp[OF Y Z]]]]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1451 |
apply (auto intro!: integrable_count_space simp: X Y Z simple_functionD) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1452 |
done |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1453 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1454 |
|
61808 | 1455 |
subsection \<open>Conditional Entropy\<close> |
39097 | 1456 |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1457 |
definition (in prob_space) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1458 |
"conditional_entropy b S T X Y = - (\<integral>(x, y). log b (enn2real (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) (x, y)) / |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1459 |
enn2real (RN_deriv T (distr M T Y) y)) \<partial>distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)))" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1460 |
|
40859 | 1461 |
abbreviation (in information_space) |
1462 |
conditional_entropy_Pow ("\<H>'(_ | _')") where |
|
47694 | 1463 |
"\<H>(X | Y) \<equiv> conditional_entropy b (count_space (X`space M)) (count_space (Y`space M)) X Y" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1464 |
|
49791 | 1465 |
lemma (in information_space) conditional_entropy_generic_eq: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1466 |
fixes Pxy :: "_ \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" |
49791 | 1467 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1468 |
assumes Py[measurable]: "distributed M T Y Py" and Py_nn[simp]: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1469 |
assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1470 |
and Pxy_nn[simp]: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1471 |
shows "conditional_entropy b S T X Y = - (\<integral>(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \<partial>(S \<Otimes>\<^sub>M T))" |
49791 | 1472 |
proof - |
1473 |
interpret S: sigma_finite_measure S by fact |
|
1474 |
interpret T: sigma_finite_measure T by fact |
|
1475 |
interpret ST: pair_sigma_finite S T .. |
|
1476 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1477 |
have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1478 |
using Py Py_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1479 |
have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1480 |
using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1481 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1482 |
have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x)). Pxy x = enn2real (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) x)" |
49791 | 1483 |
unfolding AE_density[OF distributed_borel_measurable, OF Pxy] |
1484 |
unfolding distributed_distr_eq_density[OF Pxy] |
|
1485 |
using distributed_RN_deriv[OF Pxy] |
|
1486 |
by auto |
|
1487 |
moreover |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1488 |
have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ennreal (Pxy x)). Py (snd x) = enn2real (RN_deriv T (distr M T Y) (snd x))" |
49791 | 1489 |
unfolding AE_density[OF distributed_borel_measurable, OF Pxy] |
1490 |
unfolding distributed_distr_eq_density[OF Py] |
|
1491 |
using distributed_RN_deriv[OF Py] |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1492 |
by (force intro: ST.AE_pair_measure) |
49791 | 1493 |
ultimately |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1494 |
have "conditional_entropy b S T X Y = - (\<integral>x. Pxy x * log b (Pxy x / Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))" |
49791 | 1495 |
unfolding conditional_entropy_def neg_equal_iff_equal |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1496 |
apply (subst integral_real_density[symmetric]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1497 |
apply (auto simp: distributed_distr_eq_density[OF Pxy] space_pair_measure |
49791 | 1498 |
intro!: integral_cong_AE) |
1499 |
done |
|
1500 |
then show ?thesis by (simp add: split_beta') |
|
1501 |
qed |
|
1502 |
||
1503 |
lemma (in information_space) conditional_entropy_eq_entropy: |
|
47694 | 1504 |
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" |
1505 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1506 |
assumes Py[measurable]: "distributed M T Y Py" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1507 |
and Py_nn[simp]: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1508 |
assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1509 |
and Pxy_nn[simp]: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1510 |
assumes I1: "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1511 |
assumes I2: "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1512 |
shows "conditional_entropy b S T X Y = entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) - entropy b T Y" |
40859 | 1513 |
proof - |
47694 | 1514 |
interpret S: sigma_finite_measure S by fact |
1515 |
interpret T: sigma_finite_measure T by fact |
|
1516 |
interpret ST: pair_sigma_finite S T .. |
|
1517 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1518 |
have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1519 |
using Py Py_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1520 |
have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1521 |
using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1522 |
|
47694 | 1523 |
have "entropy b T Y = - (\<integral>y. Py y * log b (Py y) \<partial>T)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1524 |
by (rule entropy_distr[OF Py Py_nn]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1525 |
also have "\<dots> = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^sub>M T))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1526 |
using b_gt_1 |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1527 |
by (subst distributed_transform_integral[OF Pxy _ Py, where T=snd]) |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1528 |
(auto intro!: Bochner_Integration.integral_cong simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1529 |
finally have e_eq: "entropy b T Y = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^sub>M T))" . |
49791 | 1530 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1531 |
have **: "\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1532 |
by (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1533 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1534 |
have ae2: "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1535 |
by (intro subdensity_real[of snd, OF _ Pxy Py]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1536 |
(auto intro: measurable_Pair simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1537 |
moreover have ae4: "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Py (snd x)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1538 |
using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'') |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1539 |
ultimately have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Pxy x \<and> 0 \<le> Py (snd x) \<and> |
49790
6b9b9ebba47d
remove unneeded assumption from conditional_entropy_generic_eq
hoelzl
parents:
49788
diff
changeset
|
1540 |
(Pxy x = 0 \<or> (Pxy x \<noteq> 0 \<longrightarrow> 0 < Pxy x \<and> 0 < Py (snd x)))" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1541 |
by (auto simp: space_pair_measure less_le) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1542 |
then have ae: "AE x in S \<Otimes>\<^sub>M T. |
47694 | 1543 |
Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1544 |
by (auto simp: log_simps field_simps b_gt_1) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1545 |
have "conditional_entropy b S T X Y = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1546 |
- (\<integral>x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1547 |
unfolding conditional_entropy_generic_eq[OF S T Py Py_nn Pxy Pxy_nn, simplified] neg_equal_iff_equal |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1548 |
using ae by (force intro: integral_cong_AE) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1549 |
also have "\<dots> = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^sub>M T)) - - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1550 |
by (simp add: Bochner_Integration.integral_diff[OF I1 I2]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1551 |
finally show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1552 |
using conditional_entropy_generic_eq[OF S T Py Py_nn Pxy Pxy_nn, simplified] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1553 |
entropy_distr[OF Pxy **, simplified] e_eq |
49791 | 1554 |
by (simp add: split_beta') |
1555 |
qed |
|
1556 |
||
1557 |
lemma (in information_space) conditional_entropy_eq_entropy_simple: |
|
1558 |
assumes X: "simple_function M X" and Y: "simple_function M Y" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1559 |
shows "\<H>(X | Y) = entropy b (count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x)) - \<H>(Y)" |
49791 | 1560 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1561 |
have "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)" |
49791 | 1562 |
(is "?P = ?C") using X Y by (simp add: simple_functionD pair_measure_count_space) |
1563 |
show ?thesis |
|
1564 |
by (rule conditional_entropy_eq_entropy sigma_finite_measure_count_space_finite |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1565 |
simple_functionD X Y simple_distributed simple_distributedI[OF _ _ refl] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1566 |
simple_distributed_joint simple_function_Pair integrable_count_space measure_nonneg)+ |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1567 |
(auto simp: \<open>?P = ?C\<close> measure_nonneg intro!: integrable_count_space simple_functionD X Y) |
39097 | 1568 |
qed |
1569 |
||
40859 | 1570 |
lemma (in information_space) conditional_entropy_eq: |
49792
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1571 |
assumes Y: "simple_distributed M Y Py" |
47694 | 1572 |
assumes XY: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy" |
1573 |
shows "\<H>(X | Y) = - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))" |
|
1574 |
proof (subst conditional_entropy_generic_eq[OF _ _ |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1575 |
simple_distributed[OF Y] _ simple_distributed_joint[OF XY]]) |
49792
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1576 |
have "finite ((\<lambda>x. (X x, Y x))`space M)" |
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1577 |
using XY unfolding simple_distributed_def by auto |
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1578 |
from finite_imageI[OF this, of fst] |
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1579 |
have [simp]: "finite (X`space M)" |
56154
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
53374
diff
changeset
|
1580 |
by (simp add: image_comp comp_def) |
47694 | 1581 |
note Y[THEN simple_distributed_finite, simp] |
1582 |
show "sigma_finite_measure (count_space (X ` space M))" |
|
1583 |
by (simp add: sigma_finite_measure_count_space_finite) |
|
1584 |
show "sigma_finite_measure (count_space (Y ` space M))" |
|
1585 |
by (simp add: sigma_finite_measure_count_space_finite) |
|
1586 |
let ?f = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x else 0)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1587 |
have "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)" |
47694 | 1588 |
(is "?P = ?C") |
49792
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1589 |
using Y by (simp add: simple_distributed_finite pair_measure_count_space) |
47694 | 1590 |
have eq: "(\<lambda>(x, y). ?f (x, y) * log b (?f (x, y) / Py y)) = |
1591 |
(\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / Py (snd x)) else 0)" |
|
1592 |
by auto |
|
49792
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1593 |
from Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) = |
47694 | 1594 |
- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))" |
64267 | 1595 |
by (auto intro!: sum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite eq sum.If_cases split_beta') |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1596 |
qed (use Y XY in auto) |
39097 | 1597 |
|
47694 | 1598 |
lemma (in information_space) conditional_mutual_information_eq_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:
41661
diff
changeset
|
1599 |
assumes X: "simple_function M X" and Y: "simple_function M Y" |
47694 | 1600 |
shows "\<I>(X ; X | Y) = \<H>(X | Y)" |
1601 |
proof - |
|
63040 | 1602 |
define Py where "Py x = (if x \<in> Y`space M then measure M (Y -` {x} \<inter> space M) else 0)" for x |
1603 |
define Pxy where "Pxy x = |
|
1604 |
(if x \<in> (\<lambda>x. (X x, Y x))`space M then measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M) else 0)" |
|
1605 |
for x |
|
1606 |
define Pxxy where "Pxxy x = |
|
1607 |
(if x \<in> (\<lambda>x. (X x, X x, Y x))`space M then measure M ((\<lambda>x. (X x, X x, Y x)) -` {x} \<inter> space M) |
|
1608 |
else 0)" |
|
1609 |
for x |
|
47694 | 1610 |
let ?M = "X`space M \<times> X`space M \<times> Y`space M" |
39097 | 1611 |
|
47694 | 1612 |
note XY = simple_function_Pair[OF X Y] |
1613 |
note XXY = simple_function_Pair[OF X XY] |
|
1614 |
have Py: "simple_distributed M Y Py" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1615 |
using Y by (rule simple_distributedI) (auto simp: Py_def measure_nonneg) |
47694 | 1616 |
have Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1617 |
using XY by (rule simple_distributedI) (auto simp: Pxy_def measure_nonneg) |
47694 | 1618 |
have Pxxy: "simple_distributed M (\<lambda>x. (X x, X x, Y x)) Pxxy" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1619 |
using XXY by (rule simple_distributedI) (auto simp: Pxxy_def measure_nonneg) |
47694 | 1620 |
have eq: "(\<lambda>x. (X x, X x, Y x)) ` space M = (\<lambda>(x, y). (x, x, y)) ` (\<lambda>x. (X x, Y x)) ` space M" |
1621 |
by auto |
|
1622 |
have inj: "\<And>A. inj_on (\<lambda>(x, y). (x, x, y)) A" |
|
1623 |
by (auto simp: inj_on_def) |
|
1624 |
have Pxxy_eq: "\<And>x y. Pxxy (x, x, y) = Pxy (x, y)" |
|
1625 |
by (auto simp: Pxxy_def Pxy_def intro!: arg_cong[where f=prob]) |
|
1626 |
have "AE x in count_space ((\<lambda>x. (X x, Y x))`space M). Py (snd x) = 0 \<longrightarrow> Pxy x = 0" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1627 |
using Py Pxy |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1628 |
by (intro subdensity_real[of snd, OF _ Pxy[THEN simple_distributed] Py[THEN simple_distributed]]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1629 |
(auto intro: measurable_Pair simp: AE_count_space) |
47694 | 1630 |
then show ?thesis |
1631 |
apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy]) |
|
49792
43f49922811d
remove unnecessary assumption from conditional_entropy_eq
hoelzl
parents:
49791
diff
changeset
|
1632 |
apply (subst conditional_entropy_eq[OF Py Pxy]) |
64267 | 1633 |
apply (auto intro!: sum.cong simp: Pxxy_eq sum_negf[symmetric] eq sum.reindex[OF inj] |
47694 | 1634 |
log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1635 |
using Py[THEN simple_distributed] Pxy[THEN simple_distributed] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1636 |
apply (auto simp add: not_le AE_count_space less_le antisym |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1637 |
simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy]) |
47694 | 1638 |
done |
1639 |
qed |
|
1640 |
||
1641 |
lemma (in information_space) conditional_entropy_nonneg: |
|
1642 |
assumes X: "simple_function M X" and Y: "simple_function M Y" shows "0 \<le> \<H>(X | Y)" |
|
1643 |
using conditional_mutual_information_eq_conditional_entropy[OF X Y] conditional_mutual_information_nonneg[OF X X Y] |
|
1644 |
by simp |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1645 |
|
61808 | 1646 |
subsection \<open>Equalities\<close> |
39097 | 1647 |
|
47694 | 1648 |
lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr: |
1649 |
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real" |
|
1650 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1651 |
assumes Px[measurable]: "distributed M S X Px" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1652 |
and Px_nn[simp]: "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1653 |
and Py[measurable]: "distributed M T Y Py" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1654 |
and Py_nn[simp]: "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1655 |
and Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1656 |
and Pxy_nn[simp]: "\<And>x y. x \<in> space S \<Longrightarrow> y \<in> space T \<Longrightarrow> 0 \<le> Pxy (x, y)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1657 |
assumes Ix: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1658 |
assumes Iy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1659 |
assumes Ixy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1660 |
shows "mutual_information b S T X Y = entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" |
40859 | 1661 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1662 |
have [measurable]: "Px \<in> S \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1663 |
using Px Px_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1664 |
have [measurable]: "Py \<in> T \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1665 |
using Py Py_nn by (intro distributed_real_measurable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1666 |
have measurable_Pxy[measurable]: "Pxy \<in> (S \<Otimes>\<^sub>M T) \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1667 |
using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1668 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1669 |
have X: "entropy b S X = - (\<integral>x. Pxy x * log b (Px (fst x)) \<partial>(S \<Otimes>\<^sub>M T))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1670 |
using b_gt_1 |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1671 |
apply (subst entropy_distr[OF Px Px_nn], simp) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1672 |
apply (subst distributed_transform_integral[OF Pxy _ Px, where T=fst]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1673 |
apply (auto intro!: integral_cong simp: space_pair_measure) |
47694 | 1674 |
done |
1675 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1676 |
have Y: "entropy b T Y = - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1677 |
using b_gt_1 |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1678 |
apply (subst entropy_distr[OF Py Py_nn], simp) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1679 |
apply (subst distributed_transform_integral[OF Pxy _ Py, where T=snd]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1680 |
apply (auto intro!: integral_cong simp: space_pair_measure) |
47694 | 1681 |
done |
1682 |
||
1683 |
interpret S: sigma_finite_measure S by fact |
|
1684 |
interpret T: sigma_finite_measure T by fact |
|
1685 |
interpret ST: pair_sigma_finite S T .. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1686 |
have ST: "sigma_finite_measure (S \<Otimes>\<^sub>M T)" .. |
47694 | 1687 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1688 |
have XY: "entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^sub>M T))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1689 |
by (subst entropy_distr[OF Pxy]) (auto intro!: integral_cong simp: space_pair_measure) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1690 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1691 |
have "AE x in S \<Otimes>\<^sub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1692 |
by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1693 |
moreover have "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1694 |
by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair simp: space_pair_measure) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1695 |
moreover have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Px (fst x)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1696 |
using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'') |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1697 |
moreover have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Py (snd x)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1698 |
using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'') |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1699 |
ultimately have "AE x in S \<Otimes>\<^sub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) = |
47694 | 1700 |
Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" |
1701 |
(is "AE x in _. ?f x = ?g x") |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1702 |
using AE_space |
47694 | 1703 |
proof eventually_elim |
60580 | 1704 |
case (elim x) |
47694 | 1705 |
show ?case |
1706 |
proof cases |
|
1707 |
assume "Pxy x \<noteq> 0" |
|
60580 | 1708 |
with elim have "0 < Px (fst x)" "0 < Py (snd x)" "0 < Pxy x" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1709 |
by (auto simp: space_pair_measure less_le) |
47694 | 1710 |
then show ?thesis |
56544 | 1711 |
using b_gt_1 by (simp add: log_simps less_imp_le field_simps) |
47694 | 1712 |
qed simp |
1713 |
qed |
|
1714 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1715 |
have "entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = integral\<^sup>L (S \<Otimes>\<^sub>M T) ?f" |
47694 | 1716 |
unfolding X Y XY |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1717 |
apply (subst Bochner_Integration.integral_diff) |
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1718 |
apply (intro Bochner_Integration.integrable_diff Ixy Ix Iy)+ |
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63626
diff
changeset
|
1719 |
apply (subst Bochner_Integration.integral_diff) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56571
diff
changeset
|
1720 |
apply (intro Ixy Ix Iy)+ |
47694 | 1721 |
apply (simp add: field_simps) |
1722 |
done |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1723 |
also have "\<dots> = integral\<^sup>L (S \<Otimes>\<^sub>M T) ?g" |
61808 | 1724 |
using \<open>AE x in _. ?f x = ?g x\<close> by (intro integral_cong_AE) auto |
47694 | 1725 |
also have "\<dots> = mutual_information b S T X Y" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1726 |
by (rule mutual_information_distr[OF S T Px _ Py _ Pxy _ , symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1727 |
(auto simp: space_pair_measure) |
47694 | 1728 |
finally show ?thesis .. |
1729 |
qed |
|
1730 |
||
49802
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1731 |
lemma (in information_space) mutual_information_eq_entropy_conditional_entropy': |
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1732 |
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real" |
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1733 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1734 |
assumes Px: "distributed M S X Px" "\<And>x. x \<in> space S \<Longrightarrow> 0 \<le> Px x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1735 |
and Py: "distributed M T Y Py" "\<And>x. x \<in> space T \<Longrightarrow> 0 \<le> Py x" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1736 |
assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1737 |
"\<And>x. x \<in> space (S \<Otimes>\<^sub>M T) \<Longrightarrow> 0 \<le> Pxy x" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1738 |
assumes Ix: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1739 |
assumes Iy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1740 |
assumes Ixy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))" |
49802
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1741 |
shows "mutual_information b S T X Y = entropy b S X - conditional_entropy b S T X Y" |
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1742 |
using |
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1743 |
mutual_information_eq_entropy_conditional_entropy_distr[OF S T Px Py Pxy Ix Iy Ixy] |
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1744 |
conditional_entropy_eq_entropy[OF S T Py Pxy Ixy Iy] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1745 |
by (simp add: space_pair_measure) |
49802
dd8dffaf84b9
continuous version of mutual_information_eq_entropy_conditional_entropy
hoelzl
parents:
49792
diff
changeset
|
1746 |
|
47694 | 1747 |
lemma (in information_space) mutual_information_eq_entropy_conditional_entropy: |
1748 |
assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y" |
|
1749 |
shows "\<I>(X ; Y) = \<H>(X) - \<H>(X | Y)" |
|
1750 |
proof - |
|
1751 |
have X: "simple_distributed M X (\<lambda>x. measure M (X -` {x} \<inter> space M))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1752 |
using sf_X by (rule simple_distributedI) (auto simp: measure_nonneg) |
47694 | 1753 |
have Y: "simple_distributed M Y (\<lambda>x. measure M (Y -` {x} \<inter> space M))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1754 |
using sf_Y by (rule simple_distributedI) (auto simp: measure_nonneg) |
47694 | 1755 |
have sf_XY: "simple_function M (\<lambda>x. (X x, Y x))" |
1756 |
using sf_X sf_Y by (rule simple_function_Pair) |
|
1757 |
then have XY: "simple_distributed M (\<lambda>x. (X x, Y x)) (\<lambda>x. measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1758 |
by (rule simple_distributedI) (auto simp: measure_nonneg) |
47694 | 1759 |
from simple_distributed_joint_finite[OF this, simp] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1760 |
have eq: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)" |
47694 | 1761 |
by (simp add: pair_measure_count_space) |
1762 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1763 |
have "\<I>(X ; Y) = \<H>(X) + \<H>(Y) - entropy b (count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1764 |
using sigma_finite_measure_count_space_finite |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1765 |
sigma_finite_measure_count_space_finite |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1766 |
simple_distributed[OF X] measure_nonneg simple_distributed[OF Y] measure_nonneg simple_distributed_joint[OF XY] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1767 |
by (rule mutual_information_eq_entropy_conditional_entropy_distr) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1768 |
(auto simp: eq integrable_count_space measure_nonneg) |
47694 | 1769 |
then show ?thesis |
49791 | 1770 |
unfolding conditional_entropy_eq_entropy_simple[OF sf_X sf_Y] by simp |
47694 | 1771 |
qed |
1772 |
||
1773 |
lemma (in information_space) mutual_information_nonneg_simple: |
|
1774 |
assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y" |
|
1775 |
shows "0 \<le> \<I>(X ; Y)" |
|
1776 |
proof - |
|
1777 |
have X: "simple_distributed M X (\<lambda>x. measure M (X -` {x} \<inter> space M))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1778 |
using sf_X by (rule simple_distributedI) (auto simp: measure_nonneg) |
47694 | 1779 |
have Y: "simple_distributed M Y (\<lambda>x. measure M (Y -` {x} \<inter> space M))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1780 |
using sf_Y by (rule simple_distributedI) (auto simp: measure_nonneg) |
47694 | 1781 |
|
1782 |
have sf_XY: "simple_function M (\<lambda>x. (X x, Y x))" |
|
1783 |
using sf_X sf_Y by (rule simple_function_Pair) |
|
1784 |
then have XY: "simple_distributed M (\<lambda>x. (X x, Y x)) (\<lambda>x. measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1785 |
by (rule simple_distributedI) (auto simp: measure_nonneg) |
47694 | 1786 |
|
1787 |
from simple_distributed_joint_finite[OF this, simp] |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1788 |
have eq: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)" |
47694 | 1789 |
by (simp add: pair_measure_count_space) |
1790 |
||
40859 | 1791 |
show ?thesis |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1792 |
by (rule mutual_information_nonneg[OF _ _ simple_distributed[OF X] _ simple_distributed[OF Y] _ simple_distributed_joint[OF XY]]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1793 |
(simp_all add: eq integrable_count_space sigma_finite_measure_count_space_finite measure_nonneg) |
40859 | 1794 |
qed |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1795 |
|
40859 | 1796 |
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:
41661
diff
changeset
|
1797 |
assumes X: "simple_function M X" and Z: "simple_function M Z" |
40859 | 1798 |
shows "\<H>(X | Z) \<le> \<H>(X)" |
36624 | 1799 |
proof - |
47694 | 1800 |
have "0 \<le> \<I>(X ; Z)" using X Z by (rule mutual_information_nonneg_simple) |
1801 |
also have "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] . |
|
1802 |
finally show ?thesis by auto |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1803 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1804 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1805 |
lemma (in information_space) |
49803 | 1806 |
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real" |
1807 |
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" |
|
1808 |
assumes Px: "finite_entropy S X Px" and Py: "finite_entropy T Y Py" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1809 |
assumes Pxy: "finite_entropy (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy" |
49803 | 1810 |
shows "conditional_entropy b S T X Y \<le> entropy b S X" |
1811 |
proof - |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1812 |
have "0 \<le> mutual_information b S T X Y" |
49803 | 1813 |
by (rule mutual_information_nonneg') fact+ |
1814 |
also have "\<dots> = entropy b S X - conditional_entropy b S T X Y" |
|
79772
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1815 |
proof (intro mutual_information_eq_entropy_conditional_entropy') |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1816 |
show "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))" |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1817 |
"integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))" |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1818 |
"integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))" |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1819 |
using assms |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1820 |
by (auto intro!: finite_entropy_integrable finite_entropy_distributed |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1821 |
finite_entropy_integrable_transform[OF Px] finite_entropy_integrable_transform[OF Py] |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1822 |
intro: finite_entropy_nn) |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1823 |
qed (use assms Px Py Pxy finite_entropy_nn finite_entropy_distributed in auto) |
49803 | 1824 |
finally show ?thesis by auto |
1825 |
qed |
|
1826 |
||
40859 | 1827 |
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:
41661
diff
changeset
|
1828 |
assumes X: "simple_function M X" and Y: "simple_function M Y" |
40859 | 1829 |
shows "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)" |
1830 |
proof - |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1831 |
note XY = simple_distributedI[OF simple_function_Pair[OF X Y] measure_nonneg refl] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1832 |
note YX = simple_distributedI[OF simple_function_Pair[OF Y X] measure_nonneg refl] |
47694 | 1833 |
note simple_distributed_joint_finite[OF this, simp] |
1834 |
let ?f = "\<lambda>x. prob ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M)" |
|
1835 |
let ?g = "\<lambda>x. prob ((\<lambda>x. (Y x, X x)) -` {x} \<inter> space M)" |
|
1836 |
let ?h = "\<lambda>x. if x \<in> (\<lambda>x. (Y x, X x)) ` space M then prob ((\<lambda>x. (Y x, X x)) -` {x} \<inter> space M) else 0" |
|
1837 |
have "\<H>(\<lambda>x. (X x, Y x)) = - (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` space M. ?f x * log b (?f x))" |
|
1838 |
using XY by (rule entropy_simple_distributed) |
|
1839 |
also have "\<dots> = - (\<Sum>x\<in>(\<lambda>(x, y). (y, x)) ` (\<lambda>x. (X x, Y x)) ` space M. ?g x * log b (?g x))" |
|
64267 | 1840 |
by (subst (2) sum.reindex) (auto simp: inj_on_def intro!: sum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"]) |
47694 | 1841 |
also have "\<dots> = - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` space M. ?h x * log b (?h x))" |
64267 | 1842 |
by (auto intro!: sum.cong) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1843 |
also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" |
49786 | 1844 |
by (subst entropy_distr[OF simple_distributed_joint[OF YX]]) |
47694 | 1845 |
(auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite |
69654 | 1846 |
cong del: sum.cong_simp intro!: sum.mono_neutral_left measure_nonneg) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50419
diff
changeset
|
1847 |
finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" . |
47694 | 1848 |
then show ?thesis |
49791 | 1849 |
unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp |
36624 | 1850 |
qed |
1851 |
||
40859 | 1852 |
lemma (in information_space) entropy_partition: |
47694 | 1853 |
assumes X: "simple_function M X" |
1854 |
shows "\<H>(X) = \<H>(f \<circ> X) + \<H>(X|f \<circ> X)" |
|
36624 | 1855 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1856 |
note fX = simple_function_compose[OF X, of f] |
47694 | 1857 |
have eq: "(\<lambda>x. ((f \<circ> X) x, X x)) ` space M = (\<lambda>x. (f x, x)) ` X ` space M" by auto |
1858 |
have inj: "\<And>A. inj_on (\<lambda>x. (f x, x)) A" |
|
1859 |
by (auto simp: inj_on_def) |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1860 |
|
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1861 |
have "\<H>(X) = - (\<Sum>x\<in>X ` space M. prob (X -` {x} \<inter> space M) * log b (prob (X -` {x} \<inter> space M)))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1862 |
by (simp add: entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1863 |
also have "\<dots> = - (\<Sum>x\<in>(\<lambda>x. ((f \<circ> X) x, X x)) ` space M. |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1864 |
prob ((\<lambda>x. ((f \<circ> X) x, X x)) -` {x} \<inter> space M) * |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1865 |
log b (prob ((\<lambda>x. ((f \<circ> X) x, X x)) -` {x} \<inter> space M)))" |
47694 | 1866 |
unfolding eq |
64267 | 1867 |
apply (subst sum.reindex[OF inj]) |
1868 |
apply (auto intro!: sum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"]) |
|
47694 | 1869 |
done |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1870 |
also have "... = \<H>(\<lambda>x. ((f \<circ> X) x, X x))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1871 |
using entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] measure_nonneg refl]] |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1872 |
by fastforce |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1873 |
also have "\<dots> = \<H>(f \<circ> X) + \<H>(X|f \<circ> X)" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1874 |
using X entropy_chain_rule by blast |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
1875 |
finally show ?thesis . |
36624 | 1876 |
qed |
1877 |
||
40859 | 1878 |
corollary (in information_space) entropy_data_processing: |
79772
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1879 |
assumes "simple_function M X" shows "\<H>(f \<circ> X) \<le> \<H>(X)" |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
79492
diff
changeset
|
1880 |
by (smt (verit) assms conditional_entropy_nonneg entropy_partition simple_function_compose) |
36624 | 1881 |
|
40859 | 1882 |
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:
41661
diff
changeset
|
1883 |
assumes X: "simple_function M X" and inj: "inj_on f (X`space M)" |
36624 | 1884 |
shows "\<H>(f \<circ> X) = \<H>(X)" |
1885 |
proof (rule antisym) |
|
40859 | 1886 |
show "\<H>(f \<circ> X) \<le> \<H>(X)" using entropy_data_processing[OF X] . |
36624 | 1887 |
next |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1888 |
have sf: "simple_function M (f \<circ> X)" |
40859 | 1889 |
using X by auto |
36624 | 1890 |
have "\<H>(X) = \<H>(the_inv_into (X`space M) f \<circ> (f \<circ> X))" |
47694 | 1891 |
unfolding o_assoc |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1892 |
apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]]) |
47694 | 1893 |
apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_compose[OF X]], where f="\<lambda>x. prob (X -` {x} \<inter> space M)"]) |
64267 | 1894 |
apply (auto intro!: sum.cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def measure_nonneg) |
47694 | 1895 |
done |
36624 | 1896 |
also have "... \<le> \<H>(f \<circ> X)" |
40859 | 1897 |
using entropy_data_processing[OF sf] . |
36624 | 1898 |
finally show "\<H>(X) \<le> \<H>(f \<circ> X)" . |
1899 |
qed |
|
1900 |
||
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1901 |
end |