| author | wenzelm | 
| Tue, 05 Nov 2024 22:05:50 +0100 | |
| changeset 81350 | 1818358373e2 | 
| parent 81179 | cf2c03967178 | 
| child 82731 | acd065f00194 | 
| permissions | -rw-r--r-- | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41413 
diff
changeset
 | 
1  | 
(* Author: Johannes Hölzl, TU München *)  | 
| 40859 | 2  | 
|
| 61808 | 3  | 
section \<open>Formalization of a Countermeasure by Koepf \& Duermuth 2009\<close>  | 
| 40859 | 4  | 
|
5  | 
theory Koepf_Duermuth_Countermeasure  | 
|
| 73392 | 6  | 
imports "HOL-Probability.Information"  | 
| 40859 | 7  | 
begin  | 
8  | 
||
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
9  | 
lemma SIGMA_image_vimage:  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
10  | 
  "snd ` (SIGMA x:f`X. f -` {x} \<inter> X) = X"
 | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
11  | 
by (auto simp: image_iff)  | 
| 40859 | 12  | 
|
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
13  | 
declare inj_split_Cons[simp]  | 
| 40859 | 14  | 
|
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
15  | 
definition extensionalD :: "'b \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'b) set" where
 | 
| 
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
16  | 
  "extensionalD d A = {f. \<forall>x. x \<notin> A \<longrightarrow> f x = d}"
 | 
| 40859 | 17  | 
|
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
18  | 
lemma extensionalD_empty[simp]: "extensionalD d {} = {\<lambda>x. d}"
 | 
| 
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
19  | 
unfolding extensionalD_def by (simp add: set_eq_iff fun_eq_iff)  | 
| 40859 | 20  | 
|
21  | 
lemma funset_eq_UN_fun_upd_I:  | 
|
22  | 
assumes *: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f(a := d) \<in> F A"  | 
|
23  | 
and **: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f a \<in> G (f(a:=d))"  | 
|
24  | 
and ***: "\<And>f x. \<lbrakk> f \<in> F A ; x \<in> G f \<rbrakk> \<Longrightarrow> f(a:=x) \<in> F (insert a A)"  | 
|
25  | 
shows "F (insert a A) = (\<Union>f\<in>F A. fun_upd f a ` (G f))"  | 
|
26  | 
proof safe  | 
|
27  | 
fix f assume f: "f \<in> F (insert a A)"  | 
|
28  | 
show "f \<in> (\<Union>f\<in>F A. fun_upd f a ` G f)"  | 
|
29  | 
proof (rule UN_I[of "f(a := d)"])  | 
|
30  | 
show "f(a := d) \<in> F A" using *[OF f] .  | 
|
31  | 
show "f \<in> fun_upd (f(a:=d)) a ` G (f(a:=d))"  | 
|
32  | 
proof (rule image_eqI[of _ _ "f a"])  | 
|
33  | 
show "f a \<in> G (f(a := d))" using **[OF f] .  | 
|
34  | 
qed simp  | 
|
35  | 
qed  | 
|
36  | 
next  | 
|
37  | 
fix f x assume "f \<in> F A" "x \<in> G f"  | 
|
38  | 
from ***[OF this] show "f(a := x) \<in> F (insert a A)" .  | 
|
39  | 
qed  | 
|
40  | 
||
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
41  | 
lemma extensionalD_insert[simp]:  | 
| 40859 | 42  | 
assumes "a \<notin> A"  | 
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
43  | 
shows "extensionalD d (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensionalD d A \<inter> (A \<rightarrow> B). (\<lambda>b. f(a := b)) ` B)"  | 
| 40859 | 44  | 
apply (rule funset_eq_UN_fun_upd_I)  | 
45  | 
using assms  | 
|
| 62390 | 46  | 
by (auto intro!: inj_onI dest: inj_onD split: if_split_asm simp: extensionalD_def)  | 
| 40859 | 47  | 
|
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
48  | 
lemma finite_extensionalD_funcset[simp, intro]:  | 
| 40859 | 49  | 
assumes "finite A" "finite B"  | 
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
50  | 
shows "finite (extensionalD d A \<inter> (A \<rightarrow> B))"  | 
| 40859 | 51  | 
using assms by induct auto  | 
52  | 
||
53  | 
lemma fun_upd_eq_iff: "f(a := b) = g(a := b') \<longleftrightarrow> b = b' \<and> f(a := d) = g(a := d)"  | 
|
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41413 
diff
changeset
 | 
54  | 
by (auto simp: fun_eq_iff)  | 
| 40859 | 55  | 
|
56  | 
lemma card_funcset:  | 
|
57  | 
assumes "finite A" "finite B"  | 
|
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
58  | 
shows "card (extensionalD d A \<inter> (A \<rightarrow> B)) = card B ^ card A"  | 
| 61808 | 59  | 
using \<open>finite A\<close> proof induct  | 
60  | 
case (insert a A) thus ?case unfolding extensionalD_insert[OF \<open>a \<notin> A\<close>]  | 
|
| 40859 | 61  | 
proof (subst card_UN_disjoint, safe, simp_all)  | 
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
62  | 
show "finite (extensionalD d A \<inter> (A \<rightarrow> B))" "\<And>f. finite (fun_upd f a ` B)"  | 
| 61808 | 63  | 
using \<open>finite B\<close> \<open>finite A\<close> by simp_all  | 
| 40859 | 64  | 
next  | 
65  | 
fix f g b b' assume "f \<noteq> g" and eq: "f(a := b) = g(a := b')" and  | 
|
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
66  | 
ext: "f \<in> extensionalD d A" "g \<in> extensionalD d A"  | 
| 40859 | 67  | 
have "f a = d" "g a = d"  | 
| 61808 | 68  | 
using ext \<open>a \<notin> A\<close> by (auto simp add: extensionalD_def)  | 
69  | 
with \<open>f \<noteq> g\<close> eq show False unfolding fun_upd_eq_iff[of _ _ b _ _ d]  | 
|
| 40859 | 70  | 
by (auto simp: fun_upd_idem fun_upd_eq_iff)  | 
71  | 
next  | 
|
| 81179 | 72  | 
have "card (fun_upd f a ` B) = card B"  | 
73  | 
if "f \<in> extensionalD d A \<inter> (A \<rightarrow> B)" for f  | 
|
74  | 
proof (auto intro!: card_image inj_onI)  | 
|
75  | 
fix b b' assume "f(a := b) = f(a := b')"  | 
|
76  | 
from fun_upd_eq_iff[THEN iffD1, OF this] show "b = b'" by simp  | 
|
77  | 
qed  | 
|
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
78  | 
then show "(\<Sum>i\<in>extensionalD d A \<inter> (A \<rightarrow> B). card (fun_upd i a ` B)) = card B * card B ^ card A"  | 
| 40859 | 79  | 
using insert by simp  | 
80  | 
qed  | 
|
81  | 
qed simp  | 
|
82  | 
||
| 64272 | 83  | 
lemma prod_sum_distrib_lists:  | 
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
84  | 
fixes P and S :: "'a set" and f :: "'a \<Rightarrow> _::semiring_0" assumes "finite S"  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
85  | 
  shows "(\<Sum>ms\<in>{ms. set ms \<subseteq> S \<and> length ms = n \<and> (\<forall>i<n. P i (ms!i))}. \<Prod>x<n. f (ms ! x)) =
 | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
86  | 
         (\<Prod>i<n. \<Sum>m\<in>{m\<in>S. P i m}. f m)"
 | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
87  | 
proof (induct n arbitrary: P)  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
88  | 
case 0 then show ?case by (simp cong: conj_cong)  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
89  | 
next  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
90  | 
case (Suc n)  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
91  | 
  have *: "{ms. set ms \<subseteq> S \<and> length ms = Suc n \<and> (\<forall>i<Suc n. P i (ms ! i))} =
 | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
92  | 
    (\<lambda>(xs, x). x#xs) ` ({ms. set ms \<subseteq> S \<and> length ms = n \<and> (\<forall>i<n. P (Suc i) (ms ! i))} \<times> {m\<in>S. P 0 m})"
 | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
93  | 
apply (auto simp: image_iff length_Suc_conv)  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
94  | 
apply force+  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
95  | 
apply (case_tac i)  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
96  | 
by force+  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
97  | 
show ?case unfolding *  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
98  | 
using Suc[of "\<lambda>i. P (Suc i)"]  | 
| 
79492
 
c1b0f64eb865
A few new results (mostly brought in from other developments)
 
paulson <lp15@cam.ac.uk> 
parents: 
73392 
diff
changeset
 | 
99  | 
by (simp add: sum.reindex sum.cartesian_product'  | 
| 64272 | 100  | 
lessThan_Suc_eq_insert_0 prod.reindex sum_distrib_right[symmetric] sum_distrib_left[symmetric])  | 
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
101  | 
qed  | 
| 
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
102  | 
|
| 47694 | 103  | 
declare space_point_measure[simp]  | 
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
104  | 
|
| 47694 | 105  | 
declare sets_point_measure[simp]  | 
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
106  | 
|
| 47694 | 107  | 
lemma measure_point_measure:  | 
108  | 
"finite \<Omega> \<Longrightarrow> A \<subseteq> \<Omega> \<Longrightarrow> (\<And>x. x \<in> \<Omega> \<Longrightarrow> 0 \<le> p x) \<Longrightarrow>  | 
|
| 62977 | 109  | 
measure (point_measure \<Omega> (\<lambda>x. ennreal (p x))) A = (\<Sum>i\<in>A. p i)"  | 
110  | 
unfolding measure_def  | 
|
| 64267 | 111  | 
by (subst emeasure_point_measure_finite) (auto simp: subset_eq sum_nonneg)  | 
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
112  | 
|
| 40859 | 113  | 
locale finite_information =  | 
114  | 
fixes \<Omega> :: "'a set"  | 
|
115  | 
and p :: "'a \<Rightarrow> real"  | 
|
116  | 
and b :: real  | 
|
117  | 
assumes finite_space[simp, intro]: "finite \<Omega>"  | 
|
118  | 
and p_sums_1[simp]: "(\<Sum>\<omega>\<in>\<Omega>. p \<omega>) = 1"  | 
|
119  | 
and positive_p[simp, intro]: "\<And>x. 0 \<le> p x"  | 
|
120  | 
and b_gt_1[simp, intro]: "1 < b"  | 
|
121  | 
||
| 64267 | 122  | 
lemma (in finite_information) positive_p_sum[simp]: "0 \<le> sum p X"  | 
123  | 
by (auto intro!: sum_nonneg)  | 
|
| 40859 | 124  | 
|
| 47694 | 125  | 
sublocale finite_information \<subseteq> prob_space "point_measure \<Omega> p"  | 
| 61169 | 126  | 
by standard (simp add: one_ereal_def emeasure_point_measure_finite)  | 
| 40859 | 127  | 
|
| 47694 | 128  | 
sublocale finite_information \<subseteq> information_space "point_measure \<Omega> p" b  | 
| 61169 | 129  | 
by standard simp  | 
| 40859 | 130  | 
|
| 64267 | 131  | 
lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> prob A = sum p A"  | 
| 47694 | 132  | 
by (auto simp: measure_point_measure)  | 
| 40859 | 133  | 
|
134  | 
locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b  | 
|
135  | 
for b :: real  | 
|
136  | 
and keys :: "'key set" and K :: "'key \<Rightarrow> real"  | 
|
137  | 
and messages :: "'message set" and M :: "'message \<Rightarrow> real" +  | 
|
138  | 
fixes observe :: "'key \<Rightarrow> 'message \<Rightarrow> 'observation"  | 
|
139  | 
and n :: nat  | 
|
140  | 
begin  | 
|
141  | 
||
142  | 
definition msgs :: "('key \<times> 'message list) set" where
 | 
|
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
143  | 
  "msgs = keys \<times> {ms. set ms \<subseteq> messages \<and> length ms = n}"
 | 
| 40859 | 144  | 
|
145  | 
definition P :: "('key \<times> 'message list) \<Rightarrow> real" where
 | 
|
| 49793 | 146  | 
"P = (\<lambda>(k, ms). K k * (\<Prod>i<n. M (ms ! i)))"  | 
| 40859 | 147  | 
|
148  | 
end  | 
|
149  | 
||
150  | 
sublocale koepf_duermuth \<subseteq> finite_information msgs P b  | 
|
| 61169 | 151  | 
proof  | 
| 40859 | 152  | 
show "finite msgs" unfolding msgs_def  | 
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
153  | 
using finite_lists_length_eq[OF M.finite_space, of n]  | 
| 40859 | 154  | 
by auto  | 
155  | 
||
156  | 
have [simp]: "\<And>A. inj_on (\<lambda>(xs, n). n # xs) A" by (force intro!: inj_onI)  | 
|
157  | 
||
| 64267 | 158  | 
note sum_distrib_left[symmetric, simp]  | 
159  | 
note sum_distrib_right[symmetric, simp]  | 
|
| 
79492
 
c1b0f64eb865
A few new results (mostly brought in from other developments)
 
paulson <lp15@cam.ac.uk> 
parents: 
73392 
diff
changeset
 | 
160  | 
note sum.cartesian_product'[simp]  | 
| 40859 | 161  | 
|
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
162  | 
have "(\<Sum>ms | set ms \<subseteq> messages \<and> length ms = n. \<Prod>x<length ms. M (ms ! x)) = 1"  | 
| 40859 | 163  | 
proof (induct n)  | 
164  | 
case 0 then show ?case by (simp cong: conj_cong)  | 
|
165  | 
next  | 
|
166  | 
case (Suc n)  | 
|
167  | 
then show ?case  | 
|
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
168  | 
by (simp add: lists_length_Suc_eq lessThan_Suc_eq_insert_0  | 
| 64272 | 169  | 
sum.reindex prod.reindex)  | 
| 40859 | 170  | 
qed  | 
| 64267 | 171  | 
then show "sum P msgs = 1"  | 
| 40859 | 172  | 
unfolding msgs_def P_def by simp  | 
| 81179 | 173  | 
have "0 \<le> (\<Prod>x\<in>A. M (f x))" for A f  | 
174  | 
by (auto simp: prod_nonneg)  | 
|
175  | 
then show "0 \<le> P x" for x  | 
|
| 40859 | 176  | 
unfolding P_def by (auto split: prod.split simp: zero_le_mult_iff)  | 
177  | 
qed auto  | 
|
178  | 
||
179  | 
context koepf_duermuth  | 
|
180  | 
begin  | 
|
181  | 
||
182  | 
definition observations :: "'observation set" where  | 
|
183  | 
"observations = (\<Union>f\<in>observe ` keys. f ` messages)"  | 
|
184  | 
||
185  | 
lemma finite_observations[simp, intro]: "finite observations"  | 
|
186  | 
unfolding observations_def by auto  | 
|
187  | 
||
188  | 
definition OB :: "'key \<times> 'message list \<Rightarrow> 'observation list" where  | 
|
189  | 
"OB = (\<lambda>(k, ms). map (observe k) ms)"  | 
|
190  | 
||
191  | 
definition t :: "'observation list \<Rightarrow> 'observation \<Rightarrow> nat" where  | 
|
| 49793 | 192  | 
  t_def2: "t seq obs = card { i. i < length seq \<and> seq ! i = obs}"
 | 
193  | 
||
| 67399 | 194  | 
lemma t_def: "t seq obs = length (filter ((=) obs) seq)"  | 
| 49793 | 195  | 
unfolding t_def2 length_filter_conv_card by (subst eq_commute) simp  | 
| 40859 | 196  | 
|
197  | 
lemma card_T_bound: "card ((t\<circ>OB)`msgs) \<le> (n+1)^card observations"  | 
|
198  | 
proof -  | 
|
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
199  | 
  have "(t\<circ>OB)`msgs \<subseteq> extensionalD 0 observations \<inter> (observations \<rightarrow> {.. n})"
 | 
| 
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
200  | 
unfolding observations_def extensionalD_def OB_def msgs_def  | 
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
201  | 
by (auto simp add: t_def comp_def image_iff subset_eq)  | 
| 
42256
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
202  | 
with finite_extensionalD_funcset  | 
| 
 
461624ffd382
Rename extensional to extensionalD (extensional is also defined in FuncSet)
 
hoelzl 
parents: 
41981 
diff
changeset
 | 
203  | 
  have "card ((t\<circ>OB)`msgs) \<le> card (extensionalD 0 observations \<inter> (observations \<rightarrow> {.. n}))"
 | 
| 40859 | 204  | 
by (rule card_mono) auto  | 
205  | 
also have "\<dots> = (n + 1) ^ card observations"  | 
|
206  | 
by (subst card_funcset) auto  | 
|
207  | 
finally show ?thesis .  | 
|
208  | 
qed  | 
|
209  | 
||
210  | 
abbreviation  | 
|
| 64267 | 211  | 
"p A \<equiv> sum P A"  | 
| 40859 | 212  | 
|
| 47694 | 213  | 
abbreviation  | 
214  | 
"\<mu> \<equiv> point_measure msgs P"  | 
|
215  | 
||
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80034 
diff
changeset
 | 
216  | 
abbreviation probability (\<open>\<P>'(_') _\<close>) where  | 
| 47694 | 217  | 
"\<P>(X) x \<equiv> measure \<mu> (X -` x \<inter> msgs)"  | 
| 40859 | 218  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80034 
diff
changeset
 | 
219  | 
abbreviation joint_probability (\<open>\<P>'(_ ; _') _\<close>) where  | 
| 47694 | 220  | 
"\<P>(X ; Y) x \<equiv> \<P>(\<lambda>x. (X x, Y x)) x"  | 
| 40859 | 221  | 
|
| 81128 | 222  | 
no_notation disj (infixr \<open>|\<close> 30)  | 
| 47694 | 223  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80034 
diff
changeset
 | 
224  | 
abbreviation conditional_probability (\<open>\<P>'(_ | _') _\<close>) where  | 
| 47694 | 225  | 
"\<P>(X | Y) x \<equiv> (\<P>(X ; Y) x) / \<P>(Y) (snd`x)"  | 
| 40859 | 226  | 
|
227  | 
notation  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80034 
diff
changeset
 | 
228  | 
entropy_Pow (\<open>\<H>'( _ ')\<close>)  | 
| 40859 | 229  | 
|
230  | 
notation  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80034 
diff
changeset
 | 
231  | 
conditional_entropy_Pow (\<open>\<H>'( _ | _ ')\<close>)  | 
| 40859 | 232  | 
|
233  | 
notation  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80034 
diff
changeset
 | 
234  | 
mutual_information_Pow (\<open>\<I>'( _ ; _ ')\<close>)  | 
| 40859 | 235  | 
|
236  | 
lemma t_eq_imp_bij_func:  | 
|
237  | 
assumes "t xs = t ys"  | 
|
238  | 
  shows "\<exists>f. bij_betw f {..<length xs} {..<length ys} \<and> (\<forall>i<length xs. xs ! i = ys ! (f i))"
 | 
|
239  | 
proof -  | 
|
| 73392 | 240  | 
from assms have \<open>mset xs = mset ys\<close>  | 
241  | 
using assms by (simp add: fun_eq_iff t_def multiset_eq_iff flip: count_mset)  | 
|
242  | 
  then obtain p where \<open>p permutes {..<length ys}\<close> \<open>permute_list p ys = xs\<close>
 | 
|
243  | 
by (rule mset_eq_permutation)  | 
|
244  | 
  then have \<open>bij_betw p {..<length xs} {..<length ys}\<close> \<open>\<forall>i<length xs. xs ! i = ys ! p i\<close>
 | 
|
245  | 
by (auto dest: permutes_imp_bij simp add: permute_list_nth)  | 
|
246  | 
then show ?thesis  | 
|
247  | 
by blast  | 
|
| 40859 | 248  | 
qed  | 
249  | 
||
250  | 
lemma \<P>_k: assumes "k \<in> keys" shows "\<P>(fst) {k} = K k"
 | 
|
251  | 
proof -  | 
|
252  | 
from assms have *:  | 
|
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
253  | 
      "fst -` {k} \<inter> msgs = {k}\<times>{ms. set ms \<subseteq> messages \<and> length ms = n}"
 | 
| 40859 | 254  | 
unfolding msgs_def by auto  | 
| 47694 | 255  | 
  show "(\<P>(fst) {k}) = K k"
 | 
256  | 
apply (simp add: \<mu>'_eq)  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41689 
diff
changeset
 | 
257  | 
apply (simp add: * P_def)  | 
| 
79492
 
c1b0f64eb865
A few new results (mostly brought in from other developments)
 
paulson <lp15@cam.ac.uk> 
parents: 
73392 
diff
changeset
 | 
258  | 
apply (simp add: sum.cartesian_product')  | 
| 64272 | 259  | 
using prod_sum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] \<open>k \<in> keys\<close>  | 
260  | 
by (auto simp add: sum_distrib_left[symmetric] subset_eq prod.neutral_const)  | 
|
| 40859 | 261  | 
qed  | 
262  | 
||
263  | 
lemma fst_image_msgs[simp]: "fst`msgs = keys"  | 
|
264  | 
proof -  | 
|
265  | 
from M.not_empty obtain m where "m \<in> messages" by auto  | 
|
| 
45715
 
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
 
hoelzl 
parents: 
45712 
diff
changeset
 | 
266  | 
  then have *: "{m. set m \<subseteq> messages \<and> length m = n} \<noteq> {}"
 | 
| 40859 | 267  | 
by (auto intro!: exI[of _ "replicate n m"])  | 
268  | 
then show ?thesis  | 
|
269  | 
unfolding msgs_def fst_image_times if_not_P[OF *] by simp  | 
|
270  | 
qed  | 
|
271  | 
||
| 64267 | 272  | 
lemma sum_distribution_cut:  | 
| 47694 | 273  | 
  "\<P>(X) {x} = (\<Sum>y \<in> Y`space \<mu>. \<P>(X ; Y) {(x, y)})"
 | 
274  | 
by (subst finite_measure_finite_Union[symmetric])  | 
|
275  | 
(auto simp add: disjoint_family_on_def inj_on_def  | 
|
276  | 
intro!: arg_cong[where f=prob])  | 
|
277  | 
||
278  | 
lemma prob_conj_imp1:  | 
|
279  | 
  "prob ({x. Q x} \<inter> msgs) = 0 \<Longrightarrow> prob ({x. Pr x \<and> Q x} \<inter> msgs) = 0"
 | 
|
280  | 
  using finite_measure_mono[of "{x. Pr x \<and> Q x} \<inter> msgs" "{x. Q x} \<inter> msgs"]
 | 
|
281  | 
  using measure_nonneg[of \<mu> "{x. Pr x \<and> Q x} \<inter> msgs"]
 | 
|
282  | 
by (simp add: subset_eq)  | 
|
283  | 
||
284  | 
lemma prob_conj_imp2:  | 
|
285  | 
  "prob ({x. Pr x} \<inter> msgs) = 0 \<Longrightarrow> prob ({x. Pr x \<and> Q x} \<inter> msgs) = 0"
 | 
|
286  | 
  using finite_measure_mono[of "{x. Pr x \<and> Q x} \<inter> msgs" "{x. Pr x} \<inter> msgs"]
 | 
|
287  | 
  using measure_nonneg[of \<mu> "{x. Pr x \<and> Q x} \<inter> msgs"]
 | 
|
288  | 
by (simp add: subset_eq)  | 
|
289  | 
||
290  | 
lemma simple_function_finite: "simple_function \<mu> f"  | 
|
291  | 
by (simp add: simple_function_def)  | 
|
292  | 
||
293  | 
lemma entropy_commute: "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))"  | 
|
| 62977 | 294  | 
apply (subst (1 2) entropy_simple_distributed[OF simple_distributedI[OF simple_function_finite _ refl]])  | 
| 47694 | 295  | 
unfolding space_point_measure  | 
296  | 
proof -  | 
|
297  | 
have eq: "(\<lambda>x. (X x, Y x)) ` msgs = (\<lambda>(x, y). (y, x)) ` (\<lambda>x. (Y x, X x)) ` msgs"  | 
|
298  | 
by auto  | 
|
299  | 
  show "- (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {x}) * log b (\<P>(X ; Y) {x})) =
 | 
|
300  | 
    - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` msgs. (\<P>(Y ; X) {x}) * log b (\<P>(Y ; X) {x}))"
 | 
|
301  | 
unfolding eq  | 
|
| 64267 | 302  | 
apply (subst sum.reindex)  | 
303  | 
apply (auto simp: vimage_def inj_on_def intro!: sum.cong arg_cong[where f="\<lambda>x. prob x * log b (prob x)"])  | 
|
| 47694 | 304  | 
done  | 
| 62977 | 305  | 
qed simp_all  | 
| 47694 | 306  | 
|
307  | 
lemma (in -) measure_eq_0I: "A = {} \<Longrightarrow> measure M A = 0" by simp
 | 
|
308  | 
||
309  | 
lemma conditional_entropy_eq_ce_with_hypothesis:  | 
|
310  | 
  "\<H>(X | Y) = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) *
 | 
|
311  | 
     log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
 | 
|
312  | 
apply (subst conditional_entropy_eq[OF  | 
|
| 62977 | 313  | 
simple_distributedI[OF simple_function_finite _ refl]  | 
314  | 
simple_distributedI[OF simple_function_finite _ refl]])  | 
|
| 47694 | 315  | 
unfolding space_point_measure  | 
316  | 
proof -  | 
|
317  | 
  have "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
 | 
|
318  | 
    - (\<Sum>x\<in>X`msgs. (\<Sum>y\<in>Y`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
 | 
|
| 64267 | 319  | 
unfolding sum.cartesian_product  | 
320  | 
apply (intro arg_cong[where f=uminus] sum.mono_neutral_left)  | 
|
| 47694 | 321  | 
apply (auto simp: vimage_def image_iff intro!: measure_eq_0I)  | 
322  | 
apply metis  | 
|
323  | 
done  | 
|
324  | 
  also have "\<dots> = - (\<Sum>y\<in>Y`msgs. (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
 | 
|
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66453 
diff
changeset
 | 
325  | 
by (subst sum.swap) rule  | 
| 47694 | 326  | 
  also have "\<dots> = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
 | 
| 64267 | 327  | 
by (auto simp add: sum_distrib_left vimage_def intro!: sum.cong prob_conj_imp1)  | 
| 47694 | 328  | 
  finally show "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
 | 
329  | 
    -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))" .
 | 
|
| 62977 | 330  | 
qed simp_all  | 
| 47694 | 331  | 
|
| 49793 | 332  | 
lemma ce_OB_eq_ce_t: "\<I>(fst ; OB) = \<I>(fst ; t\<circ>OB)"  | 
| 40859 | 333  | 
proof -  | 
| 61808 | 334  | 
txt \<open>Lemma 2\<close>  | 
| 40859 | 335  | 
|
| 81179 | 336  | 
  have t_eq_imp: "(\<P>(OB ; fst) {(obs, k)}) = (\<P>(OB ; fst) {(obs', k)})"
 | 
337  | 
if "k \<in> keys" "K k \<noteq> 0" and obs': "obs' \<in> OB ` msgs" and obs: "obs \<in> OB ` msgs"  | 
|
338  | 
and eq: "t obs = t obs'"  | 
|
339  | 
for k obs obs'  | 
|
340  | 
proof -  | 
|
341  | 
from t_eq_imp_bij_func[OF eq]  | 
|
| 40859 | 342  | 
    obtain t_f where "bij_betw t_f {..<n} {..<n}" and
 | 
343  | 
obs_t_f: "\<And>i. i<n \<Longrightarrow> obs!i = obs' ! t_f i"  | 
|
344  | 
using obs obs' unfolding OB_def msgs_def by auto  | 
|
345  | 
    then have t_f: "inj_on t_f {..<n}" "{..<n} = t_f`{..<n}" unfolding bij_betw_def by auto
 | 
|
346  | 
||
| 81179 | 347  | 
    have *: "(\<P>(OB ; fst) {(obs, k)}) / K k =
 | 
348  | 
            (\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)"
 | 
|
349  | 
if "obs \<in> OB`msgs" for obs  | 
|
350  | 
proof -  | 
|
351  | 
from that have **: "length ms = n \<Longrightarrow> OB (k, ms) = obs \<longleftrightarrow> (\<forall>i<n. observe k (ms!i) = obs ! i)"  | 
|
352  | 
for ms  | 
|
| 40859 | 353  | 
unfolding OB_def msgs_def by (simp add: image_iff list_eq_iff_nth_eq)  | 
| 47694 | 354  | 
      have "(\<P>(OB ; fst) {(obs, k)}) / K k =
 | 
| 40859 | 355  | 
          p ({k}\<times>{ms. (k,ms) \<in> msgs \<and> OB (k,ms) = obs}) / K k"
 | 
| 81179 | 356  | 
by (simp add: \<mu>'_eq) (auto intro!: arg_cong[where f=p])  | 
357  | 
      also have "\<dots> = (\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)"
 | 
|
| 61808 | 358  | 
unfolding P_def using \<open>K k \<noteq> 0\<close> \<open>k \<in> keys\<close>  | 
| 
79492
 
c1b0f64eb865
A few new results (mostly brought in from other developments)
 
paulson <lp15@cam.ac.uk> 
parents: 
73392 
diff
changeset
 | 
359  | 
apply (simp add: sum.cartesian_product' sum_divide_distrib msgs_def ** cong: conj_cong)  | 
| 64272 | 360  | 
apply (subst prod_sum_distrib_lists[OF M.finite_space]) ..  | 
| 81179 | 361  | 
finally show ?thesis .  | 
362  | 
qed  | 
|
| 40859 | 363  | 
|
| 47694 | 364  | 
    have "(\<P>(OB ; fst) {(obs, k)}) / K k = (\<P>(OB ; fst) {(obs', k)}) / K k"
 | 
| 40859 | 365  | 
unfolding *[OF obs] *[OF obs']  | 
| 64272 | 366  | 
using t_f(1) obs_t_f by (subst (2) t_f(2)) (simp add: prod.reindex)  | 
| 81179 | 367  | 
then show ?thesis using \<open>K k \<noteq> 0\<close> by auto  | 
368  | 
qed  | 
|
| 40859 | 369  | 
|
| 46731 | 370  | 
  let ?S = "\<lambda>obs. t -`{t obs} \<inter> OB`msgs"
 | 
| 81179 | 371  | 
  have P_t_eq_P_OB: "\<P>(t\<circ>OB ; fst) {(t obs, k)} = real (card (?S obs)) * \<P>(OB ; fst) {(obs, k)}"
 | 
372  | 
if "k \<in> keys" "K k \<noteq> 0" and obs: "obs \<in> OB`msgs"  | 
|
373  | 
for k obs  | 
|
374  | 
proof -  | 
|
| 40859 | 375  | 
    have *: "((\<lambda>x. (t (OB x), fst x)) -` {(t obs, k)} \<inter> msgs) =
 | 
376  | 
      (\<Union>obs'\<in>?S obs. ((\<lambda>x. (OB x, fst x)) -` {(obs', k)} \<inter> msgs))" by auto
 | 
|
377  | 
    have df: "disjoint_family_on (\<lambda>obs'. (\<lambda>x. (OB x, fst x)) -` {(obs', k)} \<inter> msgs) (?S obs)"
 | 
|
378  | 
unfolding disjoint_family_on_def by auto  | 
|
| 47694 | 379  | 
    have "\<P>(t\<circ>OB ; fst) {(t obs, k)} = (\<Sum>obs'\<in>?S obs. \<P>(OB ; fst) {(obs', k)})"
 | 
380  | 
unfolding comp_def  | 
|
| 62977 | 381  | 
using finite_measure_finite_Union[OF _ _ df]  | 
| 64267 | 382  | 
by (auto simp add: * intro!: sum_nonneg)  | 
| 47694 | 383  | 
    also have "(\<Sum>obs'\<in>?S obs. \<P>(OB ; fst) {(obs', k)}) = real (card (?S obs)) * \<P>(OB ; fst) {(obs, k)}"
 | 
| 61808 | 384  | 
by (simp add: t_eq_imp[OF \<open>k \<in> keys\<close> \<open>K k \<noteq> 0\<close> obs])  | 
| 81179 | 385  | 
finally show ?thesis .  | 
386  | 
qed  | 
|
| 40859 | 387  | 
|
| 81179 | 388  | 
  have CP_t_K: "\<P>(t\<circ>OB | fst) {(t obs, k)} =
 | 
389  | 
    real (card (t -` {t obs} \<inter> OB ` msgs)) * \<P>(OB | fst) {(obs, k)}"
 | 
|
390  | 
if k: "k \<in> keys" and obs: "obs \<in> OB`msgs" for k obs  | 
|
391  | 
using \<P>_k[OF k] P_t_eq_P_OB[OF k _ obs] by auto  | 
|
| 40859 | 392  | 
|
| 81179 | 393  | 
  have CP_T_eq_CP_O: "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(fst | OB) {(k, obs)}"
 | 
394  | 
if "k \<in> keys" and obs: "obs \<in> OB`msgs" for k obs  | 
|
395  | 
proof -  | 
|
396  | 
    from that have "t -`{t obs} \<inter> OB`msgs \<noteq> {}" (is "?S \<noteq> {}") by auto
 | 
|
| 40859 | 397  | 
then have "real (card ?S) \<noteq> 0" by auto  | 
398  | 
||
399  | 
    have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / \<P>(t\<circ>OB) {t obs}"
 | 
|
| 47694 | 400  | 
      using finite_measure_mono[of "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs" "{x. fst x = k} \<inter> msgs"]
 | 
401  | 
      using measure_nonneg[of \<mu> "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs"]
 | 
|
| 62977 | 402  | 
by (auto simp add: vimage_def conj_commute subset_eq simp del: measure_nonneg)  | 
| 47694 | 403  | 
    also have "(\<P>(t\<circ>OB) {t obs}) = (\<Sum>k'\<in>keys. (\<P>(t\<circ>OB|fst) {(t obs, k')}) * (\<P>(fst) {k'}))"
 | 
404  | 
      using finite_measure_mono[of "{x. t (OB x) = t obs \<and> fst x = k} \<inter> msgs" "{x. fst x = k} \<inter> msgs"]
 | 
|
405  | 
      using measure_nonneg[of \<mu> "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs"]
 | 
|
| 64267 | 406  | 
apply (simp add: sum_distribution_cut[of "t\<circ>OB" "t obs" fst])  | 
407  | 
apply (auto intro!: sum.cong simp: subset_eq vimage_def prob_conj_imp1)  | 
|
| 47694 | 408  | 
done  | 
| 40859 | 409  | 
    also have "\<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'}) =
 | 
410  | 
      \<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'})"
 | 
|
| 61808 | 411  | 
using CP_t_K[OF \<open>k\<in>keys\<close> obs] CP_t_K[OF _ obs] \<open>real (card ?S) \<noteq> 0\<close>  | 
| 64267 | 412  | 
by (simp only: sum_distrib_left[symmetric] ac_simps  | 
| 61808 | 413  | 
mult_divide_mult_cancel_left[OF \<open>real (card ?S) \<noteq> 0\<close>]  | 
| 64267 | 414  | 
cong: sum.cong)  | 
| 40859 | 415  | 
    also have "(\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'}) = \<P>(OB) {obs}"
 | 
| 64267 | 416  | 
using sum_distribution_cut[of OB obs fst]  | 
417  | 
by (auto intro!: sum.cong simp: prob_conj_imp1 vimage_def)  | 
|
| 40859 | 418  | 
    also have "\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / \<P>(OB) {obs} = \<P>(fst | OB) {(k, obs)}"
 | 
| 47694 | 419  | 
by (auto simp: vimage_def conj_commute prob_conj_imp2)  | 
| 81179 | 420  | 
finally show ?thesis .  | 
421  | 
qed  | 
|
| 40859 | 422  | 
|
| 46731 | 423  | 
  let ?H = "\<lambda>obs. (\<Sum>k\<in>keys. \<P>(fst|OB) {(k, obs)} * log b (\<P>(fst|OB) {(k, obs)})) :: real"
 | 
424  | 
  let ?Ht = "\<lambda>obs. (\<Sum>k\<in>keys. \<P>(fst|t\<circ>OB) {(k, obs)} * log b (\<P>(fst|t\<circ>OB) {(k, obs)})) :: real"
 | 
|
| 40859 | 425  | 
|
| 81179 | 426  | 
have *: "?H obs = ?Ht (t obs)" if obs: "obs \<in> OB`msgs" for obs  | 
427  | 
proof -  | 
|
| 40859 | 428  | 
    have "?H obs = (\<Sum>k\<in>keys. \<P>(fst|t\<circ>OB) {(k, t obs)} * log b (\<P>(fst|t\<circ>OB) {(k, t obs)}))"
 | 
429  | 
using CP_T_eq_CP_O[OF _ obs]  | 
|
430  | 
by simp  | 
|
| 81179 | 431  | 
then show ?thesis .  | 
432  | 
qed  | 
|
| 40859 | 433  | 
|
434  | 
  have **: "\<And>x f A. (\<Sum>y\<in>t-`{x}\<inter>A. f y (t y)) = (\<Sum>y\<in>t-`{x}\<inter>A. f y x)" by auto
 | 
|
435  | 
||
| 81179 | 436  | 
  have P_t_sum_P_O: "\<P>(t\<circ>OB) {t (OB x)} = (\<Sum>obs\<in>?S (OB x). \<P>(OB) {obs})" for x
 | 
437  | 
proof -  | 
|
| 40859 | 438  | 
    have *: "(\<lambda>x. t (OB x)) -` {t (OB x)} \<inter> msgs =
 | 
439  | 
      (\<Union>obs\<in>?S (OB x). OB -` {obs} \<inter> msgs)" by auto
 | 
|
440  | 
    have df: "disjoint_family_on (\<lambda>obs. OB -` {obs} \<inter> msgs) (?S (OB x))"
 | 
|
441  | 
unfolding disjoint_family_on_def by auto  | 
|
| 81179 | 442  | 
show ?thesis  | 
| 47694 | 443  | 
unfolding comp_def  | 
| 62977 | 444  | 
using finite_measure_finite_Union[OF _ _ df]  | 
| 81179 | 445  | 
by (force simp add: * intro!: sum_nonneg)  | 
446  | 
qed  | 
|
| 40859 | 447  | 
|
| 61808 | 448  | 
txt \<open>Lemma 3\<close>  | 
| 40859 | 449  | 
  have "\<H>(fst | OB) = -(\<Sum>obs\<in>OB`msgs. \<P>(OB) {obs} * ?Ht (t obs))"
 | 
| 47694 | 450  | 
unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp  | 
| 40859 | 451  | 
  also have "\<dots> = -(\<Sum>obs\<in>t`OB`msgs. \<P>(t\<circ>OB) {obs} * ?Ht obs)"
 | 
452  | 
apply (subst SIGMA_image_vimage[symmetric, of "OB`msgs" t])  | 
|
| 64267 | 453  | 
apply (subst sum.reindex)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
43920 
diff
changeset
 | 
454  | 
apply (fastforce intro!: inj_onI)  | 
| 40859 | 455  | 
apply simp  | 
| 64267 | 456  | 
apply (subst sum.Sigma[symmetric, unfolded split_def])  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
43920 
diff
changeset
 | 
457  | 
using finite_space apply fastforce  | 
| 
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
43920 
diff
changeset
 | 
458  | 
using finite_space apply fastforce  | 
| 64267 | 459  | 
apply (safe intro!: sum.cong)  | 
| 40859 | 460  | 
using P_t_sum_P_O  | 
| 64267 | 461  | 
by (simp add: sum_divide_distrib[symmetric] field_simps **  | 
462  | 
sum_distrib_left[symmetric] sum_distrib_right[symmetric])  | 
|
| 40859 | 463  | 
also have "\<dots> = \<H>(fst | t\<circ>OB)"  | 
| 47694 | 464  | 
unfolding conditional_entropy_eq_ce_with_hypothesis  | 
| 40859 | 465  | 
by (simp add: comp_def image_image[symmetric])  | 
| 49793 | 466  | 
finally show ?thesis  | 
467  | 
by (subst (1 2) mutual_information_eq_entropy_conditional_entropy) simp_all  | 
|
| 40859 | 468  | 
qed  | 
469  | 
||
470  | 
theorem "\<I>(fst ; OB) \<le> real (card observations) * log b (real n + 1)"  | 
|
471  | 
proof -  | 
|
| 49793 | 472  | 
have "\<I>(fst ; OB) = \<H>(fst) - \<H>(fst | t\<circ>OB)"  | 
473  | 
unfolding ce_OB_eq_ce_t  | 
|
474  | 
by (rule mutual_information_eq_entropy_conditional_entropy simple_function_finite)+  | 
|
| 40859 | 475  | 
also have "\<dots> = \<H>(t\<circ>OB) - \<H>(t\<circ>OB | fst)"  | 
| 70347 | 476  | 
unfolding entropy_chain_rule[symmetric, OF simple_function_finite simple_function_finite] algebra_simps  | 
| 47694 | 477  | 
by (subst entropy_commute) simp  | 
| 40859 | 478  | 
also have "\<dots> \<le> \<H>(t\<circ>OB)"  | 
| 47694 | 479  | 
using conditional_entropy_nonneg[of "t\<circ>OB" fst] by simp  | 
| 40859 | 480  | 
also have "\<dots> \<le> log b (real (card ((t\<circ>OB)`msgs)))"  | 
| 62977 | 481  | 
using entropy_le_card[of "t\<circ>OB", OF simple_distributedI[OF simple_function_finite _ refl]] by simp  | 
| 40859 | 482  | 
also have "\<dots> \<le> log b (real (n + 1)^card observations)"  | 
483  | 
using card_T_bound not_empty  | 
|
| 
80034
 
95b4fb2b5359
New material and a bit of refactoring
 
paulson <lp15@cam.ac.uk> 
parents: 
79492 
diff
changeset
 | 
484  | 
by (auto intro!: log_mono simp: card_gt_0_iff of_nat_power [symmetric] simp del: of_nat_power of_nat_Suc)  | 
| 40859 | 485  | 
also have "\<dots> = real (card observations) * log b (real n + 1)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61284 
diff
changeset
 | 
486  | 
by (simp add: log_nat_power add.commute)  | 
| 40859 | 487  | 
finally show ?thesis .  | 
488  | 
qed  | 
|
489  | 
||
490  | 
end  | 
|
491  | 
||
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41413 
diff
changeset
 | 
492  | 
end  |