author | paulson <lp15@cam.ac.uk> |
Tue, 17 Sep 2019 12:36:04 +0100 | |
changeset 70721 | 47258727fa42 |
parent 69768 | 7e4966eaf781 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
1 |
theory Dining_Cryptographers |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63040
diff
changeset
|
2 |
imports "HOL-Probability.Information" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
3 |
begin |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
4 |
|
45715
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
hoelzl
parents:
44890
diff
changeset
|
5 |
lemma image_ex1_eq: "inj_on f A \<Longrightarrow> (b \<in> f ` A) \<longleftrightarrow> (\<exists>!x \<in> A. b = f x)" |
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
hoelzl
parents:
44890
diff
changeset
|
6 |
by (unfold inj_on_def) blast |
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
hoelzl
parents:
44890
diff
changeset
|
7 |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
47694
diff
changeset
|
8 |
lemma Ex1_eq: "\<exists>!x. P x \<Longrightarrow> P x \<Longrightarrow> P y \<Longrightarrow> x = y" |
45715
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
hoelzl
parents:
44890
diff
changeset
|
9 |
by auto |
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
hoelzl
parents:
44890
diff
changeset
|
10 |
|
61808 | 11 |
subsection \<open>Define the state space\<close> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
12 |
|
61808 | 13 |
text \<open> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
14 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
15 |
We introduce the state space on which the algorithm operates. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
16 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
17 |
This contains: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
18 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
19 |
\begin{description} |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
20 |
\item[n] |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
21 |
The number of cryptographers on the table. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
22 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
23 |
\item[payer] |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
24 |
Either one of the cryptographers or the NSA. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
25 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
26 |
\item[coin] |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
27 |
The result of the coin flipping for each cryptographer. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
28 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
29 |
\item[inversion] |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
30 |
The public result for each cryptographer, e.g. the sum of the coin flipping |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
31 |
for the cryptographer, its right neighbour and the information if he paid or |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
32 |
not. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
33 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
34 |
\end{description} |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
35 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
36 |
The observables are the \emph{inversions} |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
37 |
|
61808 | 38 |
\<close> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
39 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
40 |
locale dining_cryptographers_space = |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
41 |
fixes n :: nat |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
42 |
assumes n_gt_3: "n \<ge> 3" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
43 |
begin |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
44 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
45 |
definition "dining_cryptographers = |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
46 |
({None} \<union> Some ` {0..<n}) \<times> {xs :: bool list. length xs = n}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
47 |
definition "payer dc = fst dc" |
47694 | 48 |
definition coin :: "(nat option \<times> bool list) \<Rightarrow> nat \<Rightarrow> bool" where |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
49 |
"coin dc c = snd dc ! (c mod n)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
50 |
definition "inversion dc = |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
51 |
map (\<lambda>c. (payer dc = Some c) \<noteq> (coin dc c \<noteq> coin dc (c + 1))) [0..<n]" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
52 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
53 |
definition "result dc = foldl (\<lambda> a b. a \<noteq> b) False (inversion dc)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
54 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
55 |
lemma coin_n[simp]: "coin dc n = coin dc 0" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
56 |
unfolding coin_def by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
57 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
58 |
theorem correctness: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
59 |
assumes "dc \<in> dining_cryptographers" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
60 |
shows "result dc \<longleftrightarrow> (payer dc \<noteq> None)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
61 |
proof - |
67399 | 62 |
let ?XOR = "\<lambda>f l. foldl (\<noteq>) False (map f [0..<l])" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
63 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
64 |
have foldl_coin: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
65 |
"\<not> ?XOR (\<lambda>c. coin dc c \<noteq> coin dc (c + 1)) n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
66 |
proof - |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
67 |
define n' where "n' = n" \<comment> \<open>Need to hide n, as it is hidden in coin\<close> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
68 |
have "?XOR (\<lambda>c. coin dc c \<noteq> coin dc (c + 1)) n' |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
69 |
= (coin dc 0 \<noteq> coin dc n')" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
70 |
by (induct n') auto |
61808 | 71 |
thus ?thesis using \<open>n' \<equiv> n\<close> by simp |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
72 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
73 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
74 |
from assms have "payer dc = None \<or> (\<exists>k<n. payer dc = Some k)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
75 |
unfolding dining_cryptographers_def payer_def by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
76 |
thus ?thesis |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
77 |
proof (rule disjE) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
78 |
assume "payer dc = None" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
79 |
thus ?thesis unfolding result_def inversion_def |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
80 |
using foldl_coin by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
81 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
82 |
assume "\<exists>k<n. payer dc = Some k" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
83 |
then obtain k where "k < n" and "payer dc = Some k" by auto |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
84 |
define l where "l = n" \<comment> \<open>Need to hide n, as it is hidden in coin, payer etc.\<close> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
85 |
have "?XOR (\<lambda>c. (payer dc = Some c) \<noteq> (coin dc c \<noteq> coin dc (c + 1))) l = |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
86 |
((k < l) \<noteq> ?XOR (\<lambda>c. (coin dc c \<noteq> coin dc (c + 1))) l)" |
61808 | 87 |
using \<open>payer dc = Some k\<close> by (induct l) auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
88 |
thus ?thesis |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
89 |
unfolding result_def inversion_def l_def |
61808 | 90 |
using \<open>payer dc = Some k\<close> foldl_coin \<open>k < n\<close> by simp |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
91 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
92 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
93 |
|
61808 | 94 |
text \<open> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
95 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
96 |
We now restrict the state space for the dining cryptographers to the cases when |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
97 |
one of the cryptographer pays. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
98 |
|
61808 | 99 |
\<close> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
100 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
101 |
definition |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
102 |
"dc_crypto = dining_cryptographers - {None}\<times>UNIV" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
103 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
104 |
lemma dc_crypto: "dc_crypto = Some ` {0..<n} \<times> {xs :: bool list. length xs = n}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
105 |
unfolding dc_crypto_def dining_cryptographers_def by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
106 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
107 |
lemma image_payer_dc_crypto: "payer ` dc_crypto = Some ` {0..<n}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
108 |
proof - |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
109 |
have *: "{xs. length xs = n} \<noteq> {}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
110 |
by (auto intro!: exI[of _ "replicate n undefined"]) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
111 |
show ?thesis |
46905 | 112 |
unfolding payer_def [abs_def] dc_crypto fst_image_times if_not_P[OF *] .. |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
113 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
114 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
115 |
lemma card_payer_and_inversion: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
116 |
assumes "xs \<in> inversion ` dc_crypto" and "i < n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
117 |
shows "card {dc \<in> dc_crypto. payer dc = Some i \<and> inversion dc = xs} = 2" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
118 |
(is "card ?S = 2") |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
119 |
proof - |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
120 |
obtain ys j where xs_inv: "inversion (Some j, ys) = xs" and |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
121 |
"j < n" and "(Some j, ys) \<in> dc_crypto" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
122 |
using assms(1) by (auto simp: dc_crypto) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
123 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
124 |
hence "length ys = n" by (simp add: dc_crypto) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
125 |
have [simp]: "length xs = n" using xs_inv[symmetric] by (simp add: inversion_def) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
126 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
127 |
{ fix b |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
128 |
have "inj_on (\<lambda>x. inversion (Some i, x)) {ys. ys ! 0 = b \<and> length ys = length xs}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
129 |
proof (rule inj_onI) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
130 |
fix x y |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
131 |
assume "x \<in> {ys. ys ! 0 = b \<and> length ys = length xs}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
132 |
and "y \<in> {ys. ys ! 0 = b \<and> length ys = length xs}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
133 |
and inv: "inversion (Some i, x) = inversion (Some i, y)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
134 |
hence [simp]: "x ! 0 = y ! 0" "length y = n" "length x = n" |
61808 | 135 |
using \<open>length xs = n\<close> by simp_all |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
136 |
have *: "\<And>j. j < n \<Longrightarrow> |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
137 |
(x ! j = x ! (Suc j mod n)) = (y ! j = y ! (Suc j mod n))" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
138 |
using inv unfolding inversion_def map_eq_conv payer_def coin_def |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43920
diff
changeset
|
139 |
by fastforce |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
140 |
show "x = y" |
68975
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
141 |
proof (rule nth_equalityI, simp) |
61808 | 142 |
fix j assume "j < length x" hence "j < n" using \<open>length xs = n\<close> by simp |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
143 |
thus "x ! j = y ! j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
144 |
proof (induct j) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
145 |
case (Suc j) |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
47694
diff
changeset
|
146 |
hence "j < n" by simp |
61808 | 147 |
with Suc show ?case using *[OF \<open>j < n\<close>] |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
148 |
by (cases "y ! j") simp_all |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
149 |
qed simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
150 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
151 |
qed } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
152 |
note inj_inv = this |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
153 |
|
61808 | 154 |
txt \<open> |
69597 | 155 |
We now construct the possible inversions for \<^term>\<open>xs\<close> when the payer is |
156 |
\<^term>\<open>i\<close>. |
|
61808 | 157 |
\<close> |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
158 |
|
63040 | 159 |
define zs where "zs = map (\<lambda>p. if p \<in> {min i j<..max i j} then \<not> ys ! p else ys ! p) [0..<n]" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
160 |
hence [simp]: "length zs = n" by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
161 |
hence [simp]: "0 < length zs" using n_gt_3 by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
162 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
163 |
have "\<And>l. l < max i j \<Longrightarrow> Suc l mod n = Suc l" |
61808 | 164 |
using \<open>i < n\<close> \<open>j < n\<close> by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
165 |
{ fix l assume "l < n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
166 |
hence "(((l < min i j \<or> l = min i j) \<or> (min i j < l \<and> l < max i j)) \<or> l = max i j) \<or> max i j < l" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
167 |
hence "((i = l) = (zs ! l = zs ! (Suc l mod n))) = ((j = l) = (ys ! l = ys ! (Suc l mod n)))" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
168 |
apply - proof ((erule disjE)+) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
169 |
assume "l < min i j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
170 |
hence "l \<noteq> i" and "l \<noteq> j" and "zs ! l = ys ! l" and |
61808 | 171 |
"zs ! (Suc l mod n) = ys ! (Suc l mod n)" using \<open>i < n\<close> \<open>j < n\<close> unfolding zs_def by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
172 |
thus ?thesis by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
173 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
174 |
assume "l = min i j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
175 |
show ?thesis |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
176 |
proof (cases rule: linorder_cases) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
177 |
assume "i < j" |
61808 | 178 |
hence "l = i" and "Suc l < n" and "i \<noteq> j" and "Suc l \<le> max i j" using \<open>l = min i j\<close> using \<open>j < n\<close> by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
179 |
hence "zs ! l = ys ! l" and "zs ! (Suc l mod n) = (\<not> ys ! (Suc l mod n))" |
61808 | 180 |
using \<open>l = min i j\<close>[symmetric] by (simp_all add: zs_def) |
181 |
thus ?thesis using \<open>l = i\<close> \<open>i \<noteq> j\<close> by simp |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
182 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
183 |
assume "j < i" |
61808 | 184 |
hence "l = j" and "Suc l < n" and "i \<noteq> j" and "Suc l \<le> max i j" using \<open>l = min i j\<close> using \<open>i < n\<close> by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
185 |
hence "zs ! l = ys ! l" and "zs ! (Suc l mod n) = (\<not> ys ! (Suc l mod n))" |
61808 | 186 |
using \<open>l = min i j\<close>[symmetric] by (simp_all add: zs_def) |
187 |
thus ?thesis using \<open>l = j\<close> \<open>i \<noteq> j\<close> by simp |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
188 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
189 |
assume "i = j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
190 |
hence "i = j" and "max i j = l" and "min i j = l" and "zs = ys" |
61808 | 191 |
using \<open>l = min i j\<close> by (simp_all add: zs_def \<open>length ys = n\<close>[symmetric] map_nth) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
192 |
thus ?thesis by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
193 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
194 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
195 |
assume "min i j < l \<and> l < max i j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
196 |
hence "i \<noteq> l" and "j \<noteq> l" and "zs ! l = (\<not> ys ! l)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
197 |
"zs ! (Suc l mod n) = (\<not> ys ! (Suc l mod n))" |
61808 | 198 |
using \<open>i < n\<close> \<open>j < n\<close> by (auto simp: zs_def) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
199 |
thus ?thesis by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
200 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
201 |
assume "l = max i j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
202 |
show ?thesis |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
203 |
proof (cases rule: linorder_cases) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
204 |
assume "i < j" |
61808 | 205 |
hence "l = j" and "i \<noteq> j" using \<open>l = max i j\<close> using \<open>j < n\<close> by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
206 |
have "zs ! (Suc l mod n) = ys ! (Suc l mod n)" |
61808 | 207 |
using \<open>j < n\<close> \<open>i < j\<close> \<open>l = j\<close> by (cases "Suc l = n") (auto simp add: zs_def) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
208 |
moreover have "zs ! l = (\<not> ys ! l)" |
61808 | 209 |
using \<open>j < n\<close> \<open>i < j\<close> by (auto simp add: \<open>l = j\<close> zs_def) |
210 |
ultimately show ?thesis using \<open>l = j\<close> \<open>i \<noteq> j\<close> by simp |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
211 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
212 |
assume "j < i" |
61808 | 213 |
hence "l = i" and "i \<noteq> j" using \<open>l = max i j\<close> by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
214 |
have "zs ! (Suc l mod n) = ys ! (Suc l mod n)" |
61808 | 215 |
using \<open>i < n\<close> \<open>j < i\<close> \<open>l = i\<close> by (cases "Suc l = n") (auto simp add: zs_def) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
216 |
moreover have "zs ! l = (\<not> ys ! l)" |
61808 | 217 |
using \<open>i < n\<close> \<open>j < i\<close> by (auto simp add: \<open>l = i\<close> zs_def) |
218 |
ultimately show ?thesis using \<open>l = i\<close> \<open>i \<noteq> j\<close> by auto |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
219 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
220 |
assume "i = j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
221 |
hence "i = j" and "max i j = l" and "min i j = l" and "zs = ys" |
61808 | 222 |
using \<open>l = max i j\<close> by (simp_all add: zs_def \<open>length ys = n\<close>[symmetric] map_nth) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
223 |
thus ?thesis by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
224 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
225 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
226 |
assume "max i j < l" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
227 |
hence "j \<noteq> l" and "i \<noteq> l" by simp_all |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
228 |
have "zs ! (Suc l mod n) = ys ! (Suc l mod n)" |
61808 | 229 |
using \<open>l < n\<close> \<open>max i j < l\<close> by (cases "Suc l = n") (auto simp add: zs_def) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
230 |
moreover have "zs ! l = ys ! l" |
61808 | 231 |
using \<open>l < n\<close> \<open>max i j < l\<close> by (auto simp add: zs_def) |
232 |
ultimately show ?thesis using \<open>j \<noteq> l\<close> \<open>i \<noteq> l\<close> by auto |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
233 |
qed } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
234 |
hence zs: "inversion (Some i, zs) = xs" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
235 |
by (simp add: xs_inv[symmetric] inversion_def coin_def payer_def) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
236 |
moreover |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
47694
diff
changeset
|
237 |
from zs have Not_zs: "inversion (Some i, (map Not zs)) = xs" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
238 |
by (simp add: xs_inv[symmetric] inversion_def coin_def payer_def) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
239 |
ultimately |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
240 |
have "{dc \<in> dc_crypto. payer dc = Some i \<and> inversion dc = xs} = |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
241 |
{(Some i, zs), (Some i, map Not zs)}" |
61808 | 242 |
using \<open>i < n\<close> [[ hypsubst_thin = true ]] |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
243 |
proof (safe, simp_all add:dc_crypto payer_def) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
244 |
fix b assume [simp]: "length b = n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
245 |
and *: "inversion (Some i, b) = xs" and "b \<noteq> zs" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
246 |
show "b = map Not zs" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
247 |
proof (cases "b ! 0 = zs ! 0") |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
248 |
case True |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
249 |
hence zs: "zs \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs} \<and> xs = inversion (Some i, zs)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
250 |
using zs by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
251 |
have b: "b \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs} \<and> xs = inversion (Some i, b)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
252 |
using * by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
253 |
hence "b \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs}" .. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
254 |
with *[symmetric] have "xs \<in> (\<lambda>x. inversion (Some i, x)) ` {ys. ys ! 0 = b ! 0 \<and> length ys = length xs}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
255 |
by (rule image_eqI) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
256 |
from this[unfolded image_ex1_eq[OF inj_inv]] b zs |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
257 |
have "b = zs" by (rule Ex1_eq) |
61808 | 258 |
thus ?thesis using \<open>b \<noteq> zs\<close> by simp |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
259 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
260 |
case False |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
261 |
hence zs: "map Not zs \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs} \<and> xs = inversion (Some i, map Not zs)" |
61808 | 262 |
using Not_zs by (simp add: nth_map[OF \<open>0 < length zs\<close>]) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
263 |
have b: "b \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs} \<and> xs = inversion (Some i, b)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
264 |
using * by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
265 |
hence "b \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs}" .. |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
266 |
with *[symmetric] have "xs \<in> (\<lambda>x. inversion (Some i, x)) ` {ys. ys ! 0 = b ! 0 \<and> length ys = length xs}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
267 |
by (rule image_eqI) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
268 |
from this[unfolded image_ex1_eq[OF inj_inv]] b zs |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
269 |
show "b = map Not zs" by (rule Ex1_eq) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
270 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
271 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
272 |
moreover |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
273 |
have "zs \<noteq> map Not zs" |
61808 | 274 |
using \<open>0 < length zs\<close> by (cases zs) simp_all |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
275 |
ultimately show ?thesis by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
276 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
277 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
278 |
lemma finite_dc_crypto: "finite dc_crypto" |
45715
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
hoelzl
parents:
44890
diff
changeset
|
279 |
using finite_lists_length_eq[where A="UNIV :: bool set"] |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
280 |
unfolding dc_crypto by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
281 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
282 |
lemma card_inversion: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
283 |
assumes "xs \<in> inversion ` dc_crypto" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
284 |
shows "card {dc \<in> dc_crypto. inversion dc = xs} = 2 * n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
285 |
proof - |
46731 | 286 |
let ?set = "\<lambda>i. {dc \<in> dc_crypto. payer dc = Some i \<and> inversion dc = xs}" |
287 |
let ?sets = "{?set i | i. i < n}" |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
288 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
289 |
have [simp]: "length xs = n" using assms |
46905 | 290 |
by (auto simp: dc_crypto inversion_def [abs_def]) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
291 |
|
60585 | 292 |
have "{dc \<in> dc_crypto. inversion dc = xs} = (\<Union>i < n. ?set i)" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
293 |
unfolding dc_crypto payer_def by auto |
60585 | 294 |
also have "\<dots> = (\<Union>?sets)" by auto |
295 |
finally have eq_Union: "{dc \<in> dc_crypto. inversion dc = xs} = (\<Union>?sets)" by simp |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
296 |
|
60585 | 297 |
have card_double: "2 * card ?sets = card (\<Union>?sets)" |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
298 |
proof (rule card_partition) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
299 |
show "finite ?sets" by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
300 |
{ fix i assume "i < n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
301 |
have "?set i \<subseteq> dc_crypto" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
302 |
have "finite (?set i)" using finite_dc_crypto by auto } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
303 |
thus "finite (\<Union>?sets)" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
304 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
305 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
306 |
fix c assume "c \<in> ?sets" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
307 |
thus "card c = 2" using card_payer_and_inversion[OF assms] by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
308 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
309 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
310 |
fix x y assume "x \<in> ?sets" and "y \<in> ?sets" "x \<noteq> y" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
311 |
then obtain i j where xy: "x = ?set i" "y = ?set j" by auto |
61808 | 312 |
hence "i \<noteq> j" using \<open>x \<noteq> y\<close> by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
313 |
thus "x \<inter> y = {}" using xy by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
314 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
315 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
316 |
have sets: "?sets = ?set ` {..< n}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
317 |
unfolding image_def by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
318 |
{ fix i j :: nat assume asm: "i \<noteq> j" "i < n" "j < n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
319 |
{ assume iasm: "?set i = {}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
320 |
have "card (?set i) = 2" |
61808 | 321 |
using card_payer_and_inversion[OF assms \<open>i < n\<close>] by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
322 |
hence "False" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
323 |
using iasm by auto } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
324 |
then obtain c where ci: "c \<in> ?set i" by blast |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
325 |
hence cj: "c \<notin> ?set j" using asm by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
326 |
{ assume "?set i = ?set j" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
327 |
hence "False" using ci cj by auto } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
328 |
hence "?set i \<noteq> ?set j" by auto } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
329 |
hence "inj_on ?set {..< n}" unfolding inj_on_def by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
330 |
from card_image[OF this] |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
331 |
have "card (?set ` {..< n}) = n" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
332 |
hence "card ?sets = n" using sets by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
333 |
thus ?thesis using eq_Union card_double by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
334 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
335 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
336 |
lemma card_dc_crypto: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
337 |
"card dc_crypto = n * 2^n" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
338 |
unfolding dc_crypto |
45715
efd2b952f425
use list theorems in Dining Cryptographers and Koepf Duermuth Countermeasure
hoelzl
parents:
44890
diff
changeset
|
339 |
using card_lists_length_eq[of "UNIV :: bool set"] |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
340 |
by (simp add: card_cartesian_product card_image) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
341 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
342 |
lemma card_image_inversion: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
343 |
"card (inversion ` dc_crypto) = 2^(n - 1)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
344 |
proof - |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
345 |
let ?P = "{inversion -` {x} \<inter> dc_crypto |x. x \<in> inversion ` dc_crypto}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
346 |
have "\<Union>?P = dc_crypto" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
347 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
348 |
{ fix a b assume *: "(a, b) \<in> dc_crypto" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
349 |
have inv_SOME: "inversion (SOME x. inversion x = inversion (a, b) \<and> x \<in> dc_crypto) = inversion (a, b)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
350 |
apply (rule someI2) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
351 |
by (auto simp: *) } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
352 |
note inv_SOME = this |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
353 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
354 |
{ fix a b assume *: "(a, b) \<in> dc_crypto" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
355 |
have "(SOME x. inversion x = inversion (a, b) \<and> x \<in> dc_crypto) \<in> dc_crypto" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
356 |
by (rule someI2) (auto simp: *) } |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
357 |
note SOME_inv_dc = this |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
358 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
359 |
have "bij_betw (\<lambda>s. inversion (SOME x. x \<in> s \<and> x \<in> dc_crypto)) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
360 |
{inversion -` {x} \<inter> dc_crypto |x. x \<in> inversion ` dc_crypto} |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
361 |
(inversion ` dc_crypto)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
362 |
unfolding bij_betw_def |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
363 |
by (auto intro!: inj_onI image_eqI simp: inv_SOME SOME_inv_dc) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
364 |
hence card_eq: "card {inversion -` {x} \<inter> dc_crypto |x. x \<in> inversion ` dc_crypto} = card (inversion ` dc_crypto)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
365 |
by (rule bij_betw_same_card) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
366 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
367 |
have "(2*n) * card (inversion ` dc_crypto) = card (\<Union>?P)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
368 |
unfolding card_eq[symmetric] |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
369 |
proof (rule card_partition) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
370 |
have "\<Union>?P \<subseteq> dc_crypto" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
371 |
thus "finite (\<Union>?P)" using finite_dc_crypto by (auto intro: finite_subset) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
372 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
373 |
have "?P = (\<lambda>x. inversion -` {x} \<inter> dc_crypto) ` (inversion ` dc_crypto)" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
374 |
by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
375 |
thus "finite ?P" using finite_dc_crypto by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
376 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
377 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
378 |
fix c assume "c \<in> {inversion -` {x} \<inter> dc_crypto |x. x \<in> inversion ` dc_crypto}" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
379 |
then obtain x where "c = inversion -` {x} \<inter> dc_crypto" and x: "x \<in> inversion ` dc_crypto" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
380 |
hence "c = {dc \<in> dc_crypto. inversion dc = x}" by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
381 |
thus "card c = 2 * n" using card_inversion[OF x] by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
382 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
383 |
next |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
384 |
fix x y assume "x \<in> ?P" "y \<in> ?P" and "x \<noteq> y" |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
385 |
then obtain i j where |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
386 |
x: "x = inversion -` {i} \<inter> dc_crypto" and i: "i \<in> inversion ` dc_crypto" and |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
387 |
y: "y = inversion -` {j} \<inter> dc_crypto" and j: "j \<in> inversion ` dc_crypto" by auto |
61808 | 388 |
show "x \<inter> y = {}" using x y \<open>x \<noteq> y\<close> by auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
389 |
qed |
61808 | 390 |
hence "2 * card (inversion ` dc_crypto) = 2 ^ n" unfolding \<open>\<Union>?P = dc_crypto\<close> card_dc_crypto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
391 |
using n_gt_3 by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
392 |
thus ?thesis by (cases n) auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
393 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
394 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
395 |
end |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
396 |
|
47694 | 397 |
sublocale dining_cryptographers_space \<subseteq> prob_space "uniform_count_measure dc_crypto" |
398 |
by (rule prob_space_uniform_count_measure[OF finite_dc_crypto]) |
|
399 |
(insert n_gt_3, auto simp: dc_crypto intro: exI[of _ "replicate n True"]) |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
400 |
|
47694 | 401 |
sublocale dining_cryptographers_space \<subseteq> information_space "uniform_count_measure dc_crypto" 2 |
61169 | 402 |
by standard auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
403 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
404 |
notation (in dining_cryptographers_space) |
40859 | 405 |
mutual_information_Pow ("\<I>'( _ ; _ ')") |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
406 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
407 |
notation (in dining_cryptographers_space) |
40859 | 408 |
entropy_Pow ("\<H>'( _ ')") |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
409 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
410 |
notation (in dining_cryptographers_space) |
40859 | 411 |
conditional_entropy_Pow ("\<H>'( _ | _ ')") |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
412 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
413 |
theorem (in dining_cryptographers_space) |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
414 |
"\<I>( inversion ; payer ) = 0" |
47694 | 415 |
proof (rule mutual_information_eq_0_simple) |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
416 |
have n: "0 < n" using n_gt_3 by auto |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
417 |
have card_image_inversion: |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
418 |
"real (card (inversion ` dc_crypto)) = 2^n / 2" |
61808 | 419 |
unfolding card_image_inversion using \<open>0 < n\<close> by (cases n) auto |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
420 |
|
47694 | 421 |
show inversion: "simple_distributed (uniform_count_measure dc_crypto) inversion (\<lambda>x. 2 / 2^n)" |
422 |
proof (rule simple_distributedI) |
|
423 |
show "simple_function (uniform_count_measure dc_crypto) inversion" |
|
424 |
using finite_dc_crypto |
|
425 |
by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure) |
|
426 |
fix x assume "x \<in> inversion ` space (uniform_count_measure dc_crypto)" |
|
427 |
moreover have "inversion -` {x} \<inter> dc_crypto = {dc \<in> dc_crypto. inversion dc = x}" by auto |
|
428 |
ultimately show "2 / 2^n = prob (inversion -` {x} \<inter> space (uniform_count_measure dc_crypto))" |
|
61808 | 429 |
using \<open>0 < n\<close> |
47694 | 430 |
by (simp add: card_inversion card_dc_crypto finite_dc_crypto |
431 |
subset_eq space_uniform_count_measure measure_uniform_count_measure) |
|
62977 | 432 |
qed simp |
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
433 |
|
47694 | 434 |
show "simple_distributed (uniform_count_measure dc_crypto) payer (\<lambda>x. 1 / real n)" |
435 |
proof (rule simple_distributedI) |
|
436 |
show "simple_function (uniform_count_measure dc_crypto) payer" |
|
437 |
using finite_dc_crypto |
|
438 |
by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure) |
|
439 |
fix z assume "z \<in> payer ` space (uniform_count_measure dc_crypto)" |
|
440 |
then have "payer -` {z} \<inter> dc_crypto = {z} \<times> {xs. length xs = n}" |
|
69768 | 441 |
by (auto simp: dc_crypto payer_def space_uniform_count_measure cong del: image_cong_simp) |
47694 | 442 |
hence "card (payer -` {z} \<inter> dc_crypto) = 2^n" |
443 |
using card_lists_length_eq[where A="UNIV::bool set"] |
|
444 |
by (simp add: card_cartesian_product_singleton) |
|
445 |
then show "1 / real n = prob (payer -` {z} \<inter> space (uniform_count_measure dc_crypto))" |
|
446 |
using finite_dc_crypto |
|
447 |
by (subst measure_uniform_count_measure) (auto simp add: card_dc_crypto space_uniform_count_measure) |
|
62977 | 448 |
qed simp |
47694 | 449 |
|
450 |
show "simple_distributed (uniform_count_measure dc_crypto) (\<lambda>x. (inversion x, payer x)) (\<lambda>x. 2 / (real n *2^n))" |
|
451 |
proof (rule simple_distributedI) |
|
452 |
show "simple_function (uniform_count_measure dc_crypto) (\<lambda>x. (inversion x, payer x))" |
|
453 |
using finite_dc_crypto |
|
454 |
by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure) |
|
455 |
fix x assume "x \<in> (\<lambda>x. (inversion x, payer x)) ` space (uniform_count_measure dc_crypto)" |
|
456 |
then obtain i xs where x: "x = (inversion (Some i, xs), payer (Some i, xs))" |
|
457 |
and "i < n" "length xs = n" |
|
458 |
by (simp add: image_iff space_uniform_count_measure dc_crypto Bex_def) blast |
|
459 |
then have xs: "inversion (Some i, xs) \<in> inversion`dc_crypto" and i: "Some i \<in> Some ` {0..<n}" |
|
460 |
and x: "x = (inversion (Some i, xs), Some i)" by (simp_all add: payer_def dc_crypto) |
|
63040 | 461 |
moreover define ys where "ys = inversion (Some i, xs)" |
47694 | 462 |
ultimately have ys: "ys \<in> inversion`dc_crypto" |
463 |
and "Some i \<in> Some ` {0..<n}" "x = (ys, Some i)" by simp_all |
|
464 |
then have "(\<lambda>x. (inversion x, payer x)) -` {x} \<inter> space (uniform_count_measure dc_crypto) = |
|
465 |
{dc \<in> dc_crypto. payer dc = Some (the (Some i)) \<and> inversion dc = ys}" |
|
466 |
by (auto simp add: payer_def space_uniform_count_measure) |
|
467 |
then show "2 / (real n * 2 ^ n) = prob ((\<lambda>x. (inversion x, payer x)) -` {x} \<inter> space (uniform_count_measure dc_crypto))" |
|
61808 | 468 |
using \<open>i < n\<close> ys |
47694 | 469 |
by (simp add: measure_uniform_count_measure card_payer_and_inversion finite_dc_crypto subset_eq card_dc_crypto) |
62977 | 470 |
qed simp |
47694 | 471 |
|
472 |
show "\<forall>x\<in>space (uniform_count_measure dc_crypto). 2 / (real n * 2 ^ n) = 2 / 2 ^ n * (1 / real n) " |
|
36080
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
473 |
by simp |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
474 |
qed |
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
475 |
|
0d9affa4e73c
Added Information theory and Example: dining cryptographers
hoelzl
parents:
diff
changeset
|
476 |
end |