6400
|
1 |
(* Title: HOL/Auth/Yahalom
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1996 University of Cambridge
|
|
5 |
|
|
6 |
Inductive relation "yahalom" for the Yahalom protocol.
|
|
7 |
|
11251
|
8 |
Demonstrates of why Oops is necessary. This protocol can be attacked because
|
|
9 |
it doesn't keep NB secret, but without Oops it can be "verified" anyway.
|
|
10 |
The issues are discussed in lcp's LICS 2000 invited lecture.
|
6400
|
11 |
*)
|
|
12 |
|
13956
|
13 |
header{*The Yahalom Protocol: A Flawed Version*}
|
|
14 |
|
11251
|
15 |
theory Yahalom_Bad = Shared:
|
6400
|
16 |
|
11251
|
17 |
consts yahalom :: "event list set"
|
6400
|
18 |
inductive "yahalom"
|
11251
|
19 |
intros
|
6400
|
20 |
(*Initial trace is empty*)
|
11251
|
21 |
Nil: "[] \<in> yahalom"
|
6400
|
22 |
|
|
23 |
(*The spy MAY say anything he CAN say. We do not expect him to
|
|
24 |
invent new nonces here, but he can also use NS1. Common to
|
|
25 |
all similar protocols.*)
|
11251
|
26 |
Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |]
|
|
27 |
==> Says Spy B X # evsf \<in> yahalom"
|
6400
|
28 |
|
|
29 |
(*A message that has been sent can be received by the
|
|
30 |
intended recipient.*)
|
11251
|
31 |
Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |]
|
|
32 |
==> Gets B X # evsr \<in> yahalom"
|
6400
|
33 |
|
|
34 |
(*Alice initiates a protocol run*)
|
11251
|
35 |
YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |]
|
|
36 |
==> Says A B {|Agent A, Nonce NA|} # evs1 \<in> yahalom"
|
6400
|
37 |
|
|
38 |
(*Bob's response to Alice's message.*)
|
11251
|
39 |
YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2;
|
|
40 |
Gets B {|Agent A, Nonce NA|} \<in> set evs2 |]
|
|
41 |
==> Says B Server
|
6400
|
42 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
|
11251
|
43 |
# evs2 \<in> yahalom"
|
6400
|
44 |
|
|
45 |
(*The Server receives Bob's message. He responds by sending a
|
|
46 |
new session key to Alice, with a packet for forwarding to Bob.*)
|
11251
|
47 |
YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3;
|
|
48 |
Gets Server
|
6400
|
49 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
|
11251
|
50 |
\<in> set evs3 |]
|
6400
|
51 |
==> Says Server A
|
|
52 |
{|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|},
|
|
53 |
Crypt (shrK B) {|Agent A, Key KAB|}|}
|
11251
|
54 |
# evs3 \<in> yahalom"
|
6400
|
55 |
|
|
56 |
(*Alice receives the Server's (?) message, checks her Nonce, and
|
|
57 |
uses the new session key to send Bob his Nonce. The premise
|
11251
|
58 |
A \<noteq> Server is needed to prove Says_Server_not_range.*)
|
|
59 |
YM4: "[| evs4 \<in> yahalom; A \<noteq> Server;
|
6400
|
60 |
Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
|
11251
|
61 |
\<in> set evs4;
|
|
62 |
Says A B {|Agent A, Nonce NA|} \<in> set evs4 |]
|
|
63 |
==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 \<in> yahalom"
|
|
64 |
|
|
65 |
|
|
66 |
declare Says_imp_knows_Spy [THEN analz.Inj, dest]
|
|
67 |
declare parts.Body [dest]
|
|
68 |
declare Fake_parts_insert_in_Un [dest]
|
|
69 |
declare analz_into_parts [dest]
|
|
70 |
|
|
71 |
|
|
72 |
(*A "possibility property": there are traces that reach the end*)
|
|
73 |
lemma "A \<noteq> Server
|
|
74 |
==> \<exists>X NB K. \<exists>evs \<in> yahalom.
|
|
75 |
Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs"
|
|
76 |
apply (intro exI bexI)
|
|
77 |
apply (rule_tac [2] yahalom.Nil
|
|
78 |
[THEN yahalom.YM1, THEN yahalom.Reception,
|
|
79 |
THEN yahalom.YM2, THEN yahalom.Reception,
|
|
80 |
THEN yahalom.YM3, THEN yahalom.Reception,
|
13507
|
81 |
THEN yahalom.YM4], possibility)
|
11251
|
82 |
done
|
|
83 |
|
|
84 |
lemma Gets_imp_Says:
|
|
85 |
"[| Gets B X \<in> set evs; evs \<in> yahalom |] ==> \<exists>A. Says A B X \<in> set evs"
|
|
86 |
by (erule rev_mp, erule yahalom.induct, auto)
|
|
87 |
|
|
88 |
(*Must be proved separately for each protocol*)
|
|
89 |
lemma Gets_imp_knows_Spy:
|
|
90 |
"[| Gets B X \<in> set evs; evs \<in> yahalom |] ==> X \<in> knows Spy evs"
|
|
91 |
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
|
|
92 |
|
|
93 |
declare Gets_imp_knows_Spy [THEN analz.Inj, dest]
|
|
94 |
|
|
95 |
|
|
96 |
(**** Inductive proofs about yahalom ****)
|
|
97 |
|
|
98 |
(** For reasoning about the encrypted portion of messages **)
|
|
99 |
|
|
100 |
(*Lets us treat YM4 using a similar argument as for the Fake case.*)
|
|
101 |
lemma YM4_analz_knows_Spy:
|
|
102 |
"[| Gets A {|Crypt (shrK A) Y, X|} \<in> set evs; evs \<in> yahalom |]
|
|
103 |
==> X \<in> analz (knows Spy evs)"
|
|
104 |
by blast
|
|
105 |
|
|
106 |
lemmas YM4_parts_knows_Spy =
|
|
107 |
YM4_analz_knows_Spy [THEN analz_into_parts, standard]
|
|
108 |
|
|
109 |
|
|
110 |
(** Theorems of the form X \<notin> parts (knows Spy evs) imply that NOBODY
|
|
111 |
sends messages containing X! **)
|
|
112 |
|
|
113 |
(*Spy never sees a good agent's shared key!*)
|
|
114 |
lemma Spy_see_shrK [simp]:
|
|
115 |
"evs \<in> yahalom ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
|
|
116 |
apply (erule yahalom.induct, force,
|
13507
|
117 |
drule_tac [6] YM4_parts_knows_Spy, simp_all, blast+)
|
11251
|
118 |
done
|
|
119 |
|
|
120 |
lemma Spy_analz_shrK [simp]:
|
|
121 |
"evs \<in> yahalom ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
|
|
122 |
by auto
|
|
123 |
|
|
124 |
lemma Spy_see_shrK_D [dest!]:
|
|
125 |
"[|Key (shrK A) \<in> parts (knows Spy evs); evs \<in> yahalom|] ==> A \<in> bad"
|
|
126 |
by (blast dest: Spy_see_shrK)
|
|
127 |
|
|
128 |
(*Nobody can have used non-existent keys! Needed to apply analz_insert_Key*)
|
|
129 |
lemma new_keys_not_used [rule_format, simp]:
|
|
130 |
"evs \<in> yahalom ==> Key K \<notin> used evs --> K \<notin> keysFor (parts (knows Spy evs))"
|
|
131 |
apply (erule yahalom.induct, force,
|
|
132 |
frule_tac [6] YM4_parts_knows_Spy, simp_all)
|
13926
|
133 |
txt{*Fake*}
|
|
134 |
apply (force dest!: keysFor_parts_insert)
|
|
135 |
txt{*YM3, YM4*}
|
|
136 |
apply blast+
|
11251
|
137 |
done
|
|
138 |
|
|
139 |
|
|
140 |
(****
|
|
141 |
The following is to prove theorems of the form
|
|
142 |
|
|
143 |
Key K \<in> analz (insert (Key KAB) (knows Spy evs)) ==>
|
|
144 |
Key K \<in> analz (knows Spy evs)
|
|
145 |
|
|
146 |
A more general formula must be proved inductively.
|
|
147 |
****)
|
|
148 |
|
|
149 |
(** Session keys are not used to encrypt other session keys **)
|
|
150 |
|
|
151 |
lemma analz_image_freshK [rule_format]:
|
|
152 |
"evs \<in> yahalom ==>
|
|
153 |
\<forall>K KK. KK <= - (range shrK) -->
|
|
154 |
(Key K \<in> analz (Key`KK Un (knows Spy evs))) =
|
|
155 |
(K \<in> KK | Key K \<in> analz (knows Spy evs))"
|
|
156 |
apply (erule yahalom.induct, force,
|
13507
|
157 |
drule_tac [6] YM4_analz_knows_Spy, analz_freshK, spy_analz)
|
11251
|
158 |
done
|
|
159 |
|
|
160 |
lemma analz_insert_freshK: "[| evs \<in> yahalom; KAB \<notin> range shrK |] ==>
|
11655
|
161 |
(Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =
|
11251
|
162 |
(K = KAB | Key K \<in> analz (knows Spy evs))"
|
|
163 |
by (simp only: analz_image_freshK analz_image_freshK_simps)
|
|
164 |
|
|
165 |
|
|
166 |
(*** The Key K uniquely identifies the Server's message. **)
|
|
167 |
|
|
168 |
lemma unique_session_keys:
|
|
169 |
"[| Says Server A
|
|
170 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} \<in> set evs;
|
|
171 |
Says Server A'
|
|
172 |
{|Crypt (shrK A') {|Agent B', Key K, na', nb'|}, X'|} \<in> set evs;
|
|
173 |
evs \<in> yahalom |]
|
|
174 |
==> A=A' & B=B' & na=na' & nb=nb'"
|
|
175 |
apply (erule rev_mp, erule rev_mp)
|
|
176 |
apply (erule yahalom.induct, simp_all)
|
|
177 |
(*YM3, by freshness, and YM4*)
|
|
178 |
apply blast+
|
|
179 |
done
|
|
180 |
|
|
181 |
|
|
182 |
(** Crucial secrecy property: Spy does not see the keys sent in msg YM3 **)
|
|
183 |
|
|
184 |
lemma secrecy_lemma:
|
|
185 |
"[| A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
|
|
186 |
==> Says Server A
|
|
187 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|},
|
|
188 |
Crypt (shrK B) {|Agent A, Key K|}|}
|
|
189 |
\<in> set evs -->
|
|
190 |
Key K \<notin> analz (knows Spy evs)"
|
|
191 |
apply (erule yahalom.induct, force, drule_tac [6] YM4_analz_knows_Spy)
|
13507
|
192 |
apply (simp_all add: pushes analz_insert_eq analz_insert_freshK, spy_analz) (*Fake*)
|
11251
|
193 |
apply (blast dest: unique_session_keys) (*YM3*)
|
|
194 |
done
|
|
195 |
|
|
196 |
(*Final version*)
|
|
197 |
lemma Spy_not_see_encrypted_key:
|
|
198 |
"[| Says Server A
|
|
199 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|},
|
|
200 |
Crypt (shrK B) {|Agent A, Key K|}|}
|
|
201 |
\<in> set evs;
|
|
202 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
|
|
203 |
==> Key K \<notin> analz (knows Spy evs)"
|
|
204 |
by (blast dest: secrecy_lemma)
|
|
205 |
|
6400
|
206 |
|
11251
|
207 |
(** Security Guarantee for A upon receiving YM3 **)
|
|
208 |
|
|
209 |
(*If the encrypted message appears then it originated with the Server*)
|
|
210 |
lemma A_trusts_YM3:
|
|
211 |
"[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs);
|
|
212 |
A \<notin> bad; evs \<in> yahalom |]
|
|
213 |
==> Says Server A
|
|
214 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|},
|
|
215 |
Crypt (shrK B) {|Agent A, Key K|}|}
|
|
216 |
\<in> set evs"
|
|
217 |
apply (erule rev_mp)
|
|
218 |
apply (erule yahalom.induct, force,
|
|
219 |
frule_tac [6] YM4_parts_knows_Spy, simp_all)
|
|
220 |
(*Fake, YM3*)
|
|
221 |
apply blast+
|
|
222 |
done
|
|
223 |
|
|
224 |
(*The obvious combination of A_trusts_YM3 with Spy_not_see_encrypted_key*)
|
|
225 |
lemma A_gets_good_key:
|
|
226 |
"[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs);
|
|
227 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
|
|
228 |
==> Key K \<notin> analz (knows Spy evs)"
|
|
229 |
by (blast dest!: A_trusts_YM3 Spy_not_see_encrypted_key)
|
|
230 |
|
|
231 |
(** Security Guarantees for B upon receiving YM4 **)
|
|
232 |
|
|
233 |
(*B knows, by the first part of A's message, that the Server distributed
|
|
234 |
the key for A and B. But this part says nothing about nonces.*)
|
|
235 |
lemma B_trusts_YM4_shrK:
|
|
236 |
"[| Crypt (shrK B) {|Agent A, Key K|} \<in> parts (knows Spy evs);
|
|
237 |
B \<notin> bad; evs \<in> yahalom |]
|
|
238 |
==> \<exists>NA NB. Says Server A
|
|
239 |
{|Crypt (shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|},
|
|
240 |
Crypt (shrK B) {|Agent A, Key K|}|}
|
|
241 |
\<in> set evs"
|
|
242 |
apply (erule rev_mp)
|
|
243 |
apply (erule yahalom.induct, force,
|
|
244 |
frule_tac [6] YM4_parts_knows_Spy, simp_all)
|
|
245 |
(*Fake, YM3*)
|
|
246 |
apply blast+
|
|
247 |
done
|
|
248 |
|
|
249 |
(** Up to now, the reasoning is similar to standard Yahalom. Now the
|
|
250 |
doubtful reasoning occurs. We should not be assuming that an unknown
|
|
251 |
key is secure, but the model allows us to: there is no Oops rule to
|
|
252 |
let session keys become compromised.*)
|
|
253 |
|
|
254 |
(*B knows, by the second part of A's message, that the Server distributed
|
|
255 |
the key quoting nonce NB. This part says nothing about agent names.
|
|
256 |
Secrecy of K is assumed; the valid Yahalom proof uses (and later proves)
|
|
257 |
the secrecy of NB.*)
|
|
258 |
lemma B_trusts_YM4_newK [rule_format]:
|
|
259 |
"[|Key K \<notin> analz (knows Spy evs); evs \<in> yahalom|]
|
|
260 |
==> Crypt K (Nonce NB) \<in> parts (knows Spy evs) -->
|
|
261 |
(\<exists>A B NA. Says Server A
|
|
262 |
{|Crypt (shrK A) {|Agent B, Key K,
|
|
263 |
Nonce NA, Nonce NB|},
|
|
264 |
Crypt (shrK B) {|Agent A, Key K|}|}
|
|
265 |
\<in> set evs)"
|
|
266 |
apply (erule rev_mp)
|
|
267 |
apply (erule yahalom.induct, force,
|
|
268 |
frule_tac [6] YM4_parts_knows_Spy)
|
|
269 |
apply (analz_mono_contra, simp_all)
|
|
270 |
(*Fake*)
|
|
271 |
apply blast
|
|
272 |
(*YM3*)
|
|
273 |
apply blast
|
|
274 |
(*A is uncompromised because NB is secure
|
|
275 |
A's certificate guarantees the existence of the Server message*)
|
|
276 |
apply (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad
|
|
277 |
dest: Says_imp_spies
|
|
278 |
parts.Inj [THEN parts.Fst, THEN A_trusts_YM3])
|
|
279 |
done
|
|
280 |
|
|
281 |
|
|
282 |
(*B's session key guarantee from YM4. The two certificates contribute to a
|
|
283 |
single conclusion about the Server's message. *)
|
|
284 |
lemma B_trusts_YM4:
|
|
285 |
"[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},
|
|
286 |
Crypt K (Nonce NB)|} \<in> set evs;
|
|
287 |
Says B Server
|
|
288 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
|
|
289 |
\<in> set evs;
|
|
290 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
|
|
291 |
==> \<exists>na nb. Says Server A
|
|
292 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|},
|
|
293 |
Crypt (shrK B) {|Agent A, Key K|}|}
|
|
294 |
\<in> set evs"
|
|
295 |
by (blast dest: B_trusts_YM4_newK B_trusts_YM4_shrK Spy_not_see_encrypted_key
|
|
296 |
unique_session_keys)
|
|
297 |
|
|
298 |
|
|
299 |
(*The obvious combination of B_trusts_YM4 with Spy_not_see_encrypted_key*)
|
|
300 |
lemma B_gets_good_key:
|
|
301 |
"[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},
|
|
302 |
Crypt K (Nonce NB)|} \<in> set evs;
|
|
303 |
Says B Server
|
|
304 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
|
|
305 |
\<in> set evs;
|
|
306 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
|
|
307 |
==> Key K \<notin> analz (knows Spy evs)"
|
|
308 |
by (blast dest!: B_trusts_YM4 Spy_not_see_encrypted_key)
|
|
309 |
|
|
310 |
|
|
311 |
(*** Authenticating B to A: these proofs are not considered.
|
|
312 |
They are irrelevant to showing the need for Oops. ***)
|
|
313 |
|
|
314 |
|
|
315 |
(*** Authenticating A to B using the certificate Crypt K (Nonce NB) ***)
|
|
316 |
|
|
317 |
(*Assuming the session key is secure, if both certificates are present then
|
|
318 |
A has said NB. We can't be sure about the rest of A's message, but only
|
|
319 |
NB matters for freshness.*)
|
|
320 |
lemma A_Said_YM3_lemma [rule_format]:
|
|
321 |
"evs \<in> yahalom
|
|
322 |
==> Key K \<notin> analz (knows Spy evs) -->
|
|
323 |
Crypt K (Nonce NB) \<in> parts (knows Spy evs) -->
|
|
324 |
Crypt (shrK B) {|Agent A, Key K|} \<in> parts (knows Spy evs) -->
|
|
325 |
B \<notin> bad -->
|
|
326 |
(\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs)"
|
|
327 |
apply (erule yahalom.induct, force,
|
|
328 |
frule_tac [6] YM4_parts_knows_Spy)
|
|
329 |
apply (analz_mono_contra, simp_all)
|
|
330 |
(*Fake*)
|
|
331 |
apply blast
|
|
332 |
(*YM3: by new_keys_not_used we note that Crypt K (Nonce NB) could not exist*)
|
|
333 |
apply (force dest!: Crypt_imp_keysFor)
|
|
334 |
(*YM4: was Crypt K (Nonce NB) the very last message? If not, use ind. hyp.*)
|
|
335 |
apply (simp add: ex_disj_distrib)
|
|
336 |
(*yes: apply unicity of session keys*)
|
|
337 |
apply (blast dest!: Gets_imp_Says A_trusts_YM3 B_trusts_YM4_shrK
|
|
338 |
Crypt_Spy_analz_bad
|
|
339 |
dest: Says_imp_knows_Spy [THEN parts.Inj] unique_session_keys)
|
|
340 |
done
|
|
341 |
|
|
342 |
(*If B receives YM4 then A has used nonce NB (and therefore is alive).
|
|
343 |
Moreover, A associates K with NB (thus is talking about the same run).
|
|
344 |
Other premises guarantee secrecy of K.*)
|
|
345 |
lemma YM4_imp_A_Said_YM3 [rule_format]:
|
|
346 |
"[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},
|
|
347 |
Crypt K (Nonce NB)|} \<in> set evs;
|
|
348 |
Says B Server
|
|
349 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
|
|
350 |
\<in> set evs;
|
|
351 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
|
|
352 |
==> \<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs"
|
|
353 |
apply (blast intro!: A_Said_YM3_lemma
|
|
354 |
dest: Spy_not_see_encrypted_key B_trusts_YM4 Gets_imp_Says)
|
|
355 |
done
|
6400
|
356 |
|
|
357 |
end
|