author | paulson |
Mon, 05 Mar 2001 15:25:11 +0100 | |
changeset 11193 | 851c90b23a9e |
parent 11188 | 5d539f1682c3 |
child 11251 | a6816d47f41d |
permissions | -rw-r--r-- |
1934 | 1 |
(* Title: HOL/Auth/NS_Shared |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
Inductive relation "ns_shared" for Needham-Schroeder Shared-Key protocol. |
|
7 |
||
8 |
From page 247 of |
|
9 |
Burrows, Abadi and Needham. A Logic of Authentication. |
|
10 |
Proc. Royal Soc. 426 (1989) |
|
11 |
*) |
|
12 |
||
11104 | 13 |
theory NS_Shared = Shared: |
1934 | 14 |
|
11104 | 15 |
consts ns_shared :: "event list set" |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3465
diff
changeset
|
16 |
inductive "ns_shared" |
11104 | 17 |
intros |
18 |
(*Initial trace is empty*) |
|
19 |
Nil: "[] \<in> ns_shared" |
|
20 |
(*The spy MAY say anything he CAN say. We do not expect him to |
|
21 |
invent new nonces here, but he can also use NS1. Common to |
|
22 |
all similar protocols.*) |
|
23 |
Fake: "\<lbrakk>evs \<in> ns_shared; X \<in> synth (analz (spies evs))\<rbrakk> |
|
24 |
\<Longrightarrow> Says Spy B X # evs \<in> ns_shared" |
|
25 |
||
26 |
(*Alice initiates a protocol run, requesting to talk to any B*) |
|
27 |
NS1: "\<lbrakk>evs1 \<in> ns_shared; Nonce NA \<notin> used evs1\<rbrakk> |
|
28 |
\<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared" |
|
29 |
||
30 |
(*Server's response to Alice's message. |
|
31 |
!! It may respond more than once to A's request !! |
|
32 |
Server doesn't know who the true sender is, hence the A' in |
|
33 |
the sender field.*) |
|
34 |
NS2: "\<lbrakk>evs2 \<in> ns_shared; Key KAB \<notin> used evs2; |
|
35 |
Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk> |
|
36 |
\<Longrightarrow> Says Server A |
|
37 |
(Crypt (shrK A) |
|
38 |
\<lbrace>Nonce NA, Agent B, Key KAB, |
|
39 |
(Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>) |
|
40 |
# evs2 \<in> ns_shared" |
|
41 |
||
42 |
(*We can't assume S=Server. Agent A "remembers" her nonce. |
|
43 |
Need A \<noteq> Server because we allow messages to self.*) |
|
44 |
NS3: "\<lbrakk>evs3 \<in> ns_shared; A \<noteq> Server; |
|
45 |
Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3; |
|
46 |
Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk> |
|
47 |
\<Longrightarrow> Says A B X # evs3 \<in> ns_shared" |
|
48 |
||
49 |
(*Bob's nonce exchange. He does not know who the message came |
|
50 |
from, but responds to A because she is mentioned inside.*) |
|
51 |
NS4: "\<lbrakk>evs4 \<in> ns_shared; Nonce NB \<notin> used evs4; |
|
52 |
Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk> |
|
53 |
\<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared" |
|
1934 | 54 |
|
11104 | 55 |
(*Alice responds with Nonce NB if she has seen the key before. |
56 |
Maybe should somehow check Nonce NA again. |
|
57 |
We do NOT send NB-1 or similar as the Spy cannot spoof such things. |
|
58 |
Letting the Spy add or subtract 1 lets him send \<forall>nonces. |
|
59 |
Instead we distinguish the messages by sending the nonce twice.*) |
|
60 |
NS5: "\<lbrakk>evs5 \<in> ns_shared; |
|
61 |
Says B' A (Crypt K (Nonce NB)) \<in> set evs5; |
|
62 |
Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) |
|
63 |
\<in> set evs5\<rbrakk> |
|
64 |
\<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared" |
|
65 |
||
66 |
(*This message models possible leaks of session keys. |
|
67 |
The two Nonces identify the protocol run: the rule insists upon |
|
68 |
the true senders in order to make them accurate.*) |
|
69 |
Oops: "\<lbrakk>evso \<in> ns_shared; Says B A (Crypt K (Nonce NB)) \<in> set evso; |
|
70 |
Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) |
|
71 |
\<in> set evso\<rbrakk> |
|
72 |
\<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared" |
|
73 |
||
11150 | 74 |
|
75 |
declare Says_imp_knows_Spy [THEN parts.Inj, dest] |
|
76 |
declare parts.Body [dest] |
|
77 |
declare MPair_parts [elim!] (*can speed up some proofs*) |
|
78 |
||
11104 | 79 |
declare analz_subset_parts [THEN subsetD, dest] |
11150 | 80 |
declare Fake_parts_insert [THEN subsetD, dest] |
11104 | 81 |
declare image_eq_UN [simp] (*accelerates proofs involving nested images*) |
82 |
||
83 |
||
84 |
(*A "possibility property": there are traces that reach the end*) |
|
85 |
lemma "A \<noteq> Server \<Longrightarrow> \<exists>N K. \<exists>evs \<in> ns_shared. |
|
86 |
Says A B (Crypt K \<lbrace>Nonce N, Nonce N\<rbrace>) \<in> set evs" |
|
87 |
apply (intro exI bexI) |
|
88 |
apply (rule_tac [2] ns_shared.Nil |
|
89 |
[THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3, |
|
90 |
THEN ns_shared.NS4, THEN ns_shared.NS5]) |
|
91 |
apply possibility |
|
92 |
done |
|
93 |
||
94 |
(*This version is similar, while instantiating ?K and ?N to epsilon-terms |
|
95 |
lemma "A \<noteq> Server \<Longrightarrow> \<exists>evs \<in> ns_shared. |
|
96 |
Says A B (Crypt ?K \<lbrace>Nonce ?N, Nonce ?N\<rbrace>) \<in> set evs" |
|
97 |
*) |
|
98 |
||
99 |
||
100 |
(**** Inductive proofs about ns_shared ****) |
|
101 |
||
102 |
(*Forwarding lemmas, to aid simplification*) |
|
1934 | 103 |
|
11104 | 104 |
(*For reasoning about the encrypted portion of message NS3*) |
105 |
lemma NS3_msg_in_parts_spies: |
|
106 |
"Says S A (Crypt KA \<lbrace>N, B, K, X\<rbrace>) \<in> set evs \<Longrightarrow> X \<in> parts (spies evs)" |
|
107 |
by blast |
|
108 |
||
109 |
(*For reasoning about the Oops message*) |
|
110 |
lemma Oops_parts_spies: |
|
111 |
"Says Server A (Crypt (shrK A) \<lbrace>NA, B, K, X\<rbrace>) \<in> set evs |
|
112 |
\<Longrightarrow> K \<in> parts (spies evs)" |
|
113 |
by blast |
|
114 |
||
115 |
(** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY |
|
116 |
sends messages containing X! **) |
|
117 |
||
118 |
(*Spy never sees another agent's shared key! (unless it's bad at start)*) |
|
119 |
lemma Spy_see_shrK [simp]: |
|
120 |
"evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" |
|
121 |
apply (erule ns_shared.induct, force, frule_tac [4] NS3_msg_in_parts_spies) |
|
122 |
apply simp_all |
|
123 |
apply blast+; |
|
124 |
done |
|
125 |
||
126 |
||
127 |
lemma Spy_analz_shrK [simp]: |
|
128 |
"evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)" |
|
129 |
by auto |
|
130 |
||
131 |
||
132 |
(*Nobody can have used non-existent keys!*) |
|
133 |
lemma new_keys_not_used [rule_format, simp]: |
|
134 |
"evs \<in> ns_shared \<Longrightarrow> Key K \<notin> used evs \<longrightarrow> K \<notin> keysFor (parts (spies evs))" |
|
135 |
apply (erule ns_shared.induct, force, frule_tac [4] NS3_msg_in_parts_spies) |
|
136 |
apply simp_all |
|
137 |
(*Fake, NS2, NS4, NS5*) |
|
138 |
apply (blast dest!: keysFor_parts_insert)+ |
|
139 |
done |
|
140 |
||
141 |
||
142 |
(** Lemmas concerning the form of items passed in messages **) |
|
143 |
||
144 |
(*Describes the form of K, X and K' when the Server sends this message.*) |
|
145 |
lemma Says_Server_message_form: |
|
146 |
"\<lbrakk>Says Server A (Crypt K' \<lbrace>N, Agent B, Key K, X\<rbrace>) \<in> set evs; |
|
147 |
evs \<in> ns_shared\<rbrakk> |
|
148 |
\<Longrightarrow> K \<notin> range shrK \<and> |
|
149 |
X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<and> |
|
150 |
K' = shrK A" |
|
151 |
by (erule rev_mp, erule ns_shared.induct, auto) |
|
152 |
||
1934 | 153 |
|
11104 | 154 |
(*If the encrypted message appears then it originated with the Server*) |
155 |
lemma A_trusts_NS2: |
|
156 |
"\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); |
|
157 |
A \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
158 |
\<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs" |
|
159 |
apply (erule rev_mp) |
|
160 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies) |
|
161 |
apply auto |
|
162 |
done |
|
163 |
||
164 |
lemma cert_A_form: |
|
165 |
"\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); |
|
166 |
A \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
167 |
\<Longrightarrow> K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)" |
|
168 |
by (blast dest!: A_trusts_NS2 Says_Server_message_form) |
|
169 |
||
170 |
(*EITHER describes the form of X when the following message is sent, |
|
171 |
OR reduces it to the Fake case. |
|
172 |
Use Says_Server_message_form if applicable.*) |
|
173 |
lemma Says_S_message_form: |
|
174 |
"\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
|
175 |
evs \<in> ns_shared\<rbrakk> |
|
176 |
\<Longrightarrow> (K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)) |
|
177 |
\<or> X \<in> analz (spies evs)" |
|
11150 | 178 |
by (blast dest: Says_imp_knows_Spy cert_A_form analz.Inj) |
179 |
||
11104 | 180 |
|
181 |
(*Alternative version also provable |
|
182 |
lemma Says_S_message_form2: |
|
183 |
"\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
|
184 |
evs \<in> ns_shared\<rbrakk> |
|
185 |
\<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs |
|
186 |
\<or> X \<in> analz (spies evs)" |
|
187 |
apply (case_tac "A \<in> bad") |
|
188 |
apply (force dest!: Says_imp_knows_Spy [THEN analz.Inj]); |
|
189 |
by (blast dest!: A_trusts_NS2 Says_Server_message_form) |
|
190 |
*) |
|
191 |
||
192 |
||
193 |
(**** |
|
194 |
SESSION KEY COMPROMISE THEOREM. To prove theorems of the form |
|
195 |
||
196 |
Key K \<in> analz (insert (Key KAB) (spies evs)) \<Longrightarrow> |
|
197 |
Key K \<in> analz (spies evs) |
|
198 |
||
199 |
A more general formula must be proved inductively. |
|
200 |
****) |
|
1934 | 201 |
|
11104 | 202 |
(*NOT useful in this form, but it says that session keys are not used |
203 |
to encrypt messages containing other keys, in the actual protocol. |
|
204 |
We require that agents should behave like this subsequently also.*) |
|
205 |
lemma "\<lbrakk>evs \<in> ns_shared; Kab \<notin> range shrK\<rbrakk> \<Longrightarrow> |
|
206 |
(Crypt KAB X) \<in> parts (spies evs) \<and> |
|
207 |
Key K \<in> parts {X} \<longrightarrow> Key K \<in> parts (spies evs)" |
|
208 |
apply (erule ns_shared.induct, force, frule_tac [4] NS3_msg_in_parts_spies) |
|
209 |
apply simp_all |
|
210 |
(*Fake*) |
|
211 |
apply (blast dest: parts_insert_subset_Un) |
|
212 |
(*Base, NS4 and NS5*) |
|
213 |
apply auto |
|
214 |
done |
|
215 |
||
216 |
||
217 |
(** Session keys are not used to encrypt other session keys **) |
|
218 |
||
219 |
(*The equality makes the induction hypothesis easier to apply*) |
|
220 |
||
221 |
lemma analz_image_freshK [rule_format]: |
|
222 |
"evs \<in> ns_shared \<Longrightarrow> |
|
223 |
\<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow> |
|
224 |
(Key K \<in> analz (Key`KK \<union> (spies evs))) = |
|
225 |
(K \<in> KK \<or> Key K \<in> analz (spies evs))" |
|
226 |
apply (erule ns_shared.induct, force) |
|
227 |
apply (frule_tac [7] Says_Server_message_form) |
|
228 |
apply (erule_tac [4] Says_S_message_form [THEN disjE]) |
|
11117 | 229 |
apply analz_freshK |
230 |
apply spy_analz |
|
11104 | 231 |
done |
232 |
||
233 |
||
234 |
lemma analz_insert_freshK: |
|
235 |
"\<lbrakk>evs \<in> ns_shared; KAB \<notin> range shrK\<rbrakk> \<Longrightarrow> |
|
236 |
Key K \<in> analz (insert (Key KAB) (spies evs)) = |
|
237 |
(K = KAB \<or> Key K \<in> analz (spies evs))" |
|
238 |
by (simp only: analz_image_freshK analz_image_freshK_simps) |
|
239 |
||
240 |
||
241 |
(** The session key K uniquely identifies the message **) |
|
1934 | 242 |
|
11104 | 243 |
(*In messages of this form, the session key uniquely identifies the rest*) |
244 |
lemma unique_session_keys: |
|
245 |
"\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
|
246 |
Says Server A' (Crypt (shrK A') \<lbrace>NA', Agent B', Key K, X'\<rbrace>) \<in> set evs; |
|
247 |
evs \<in> ns_shared\<rbrakk> \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B' \<and> X = X'" |
|
248 |
apply (erule rev_mp, erule rev_mp, erule ns_shared.induct) |
|
249 |
apply simp_all |
|
250 |
apply blast+ |
|
251 |
done |
|
252 |
||
253 |
||
254 |
(** Crucial secrecy property: Spy does not see the keys sent in msg NS2 **) |
|
255 |
||
11150 | 256 |
(*Beware of [rule_format] and the universal quantifier!*) |
257 |
lemma secrecy_lemma: |
|
11104 | 258 |
"\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, |
259 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) |
|
260 |
\<in> set evs; |
|
261 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
262 |
\<Longrightarrow> (\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs) \<longrightarrow> |
|
263 |
Key K \<notin> analz (spies evs)" |
|
264 |
apply (erule rev_mp) |
|
265 |
apply (erule ns_shared.induct, force) |
|
266 |
apply (frule_tac [7] Says_Server_message_form) |
|
267 |
apply (frule_tac [4] Says_S_message_form) |
|
268 |
apply (erule_tac [5] disjE) |
|
269 |
apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs) |
|
270 |
apply spy_analz (*Fake*) |
|
271 |
apply blast (*NS2*) |
|
11188 | 272 |
(*NS3, Server sub-case*) (**LEVEL 8 **) |
273 |
apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2 |
|
274 |
dest: Says_imp_knows_Spy analz.Inj unique_session_keys) |
|
275 |
(*NS3, Spy sub-case; also Oops*) |
|
276 |
apply (blast dest: unique_session_keys)+ |
|
11104 | 277 |
done |
278 |
||
279 |
||
11188 | 280 |
|
11104 | 281 |
(*Final version: Server's message in the most abstract form*) |
282 |
lemma Spy_not_see_encrypted_key: |
|
283 |
"\<lbrakk>Says Server A (Crypt K' \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
|
284 |
\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; |
|
285 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
286 |
\<Longrightarrow> Key K \<notin> analz (spies evs)" |
|
11150 | 287 |
by (blast dest: Says_Server_message_form secrecy_lemma) |
11104 | 288 |
|
289 |
||
290 |
(**** Guarantees available at various stages of protocol ***) |
|
1934 | 291 |
|
11104 | 292 |
(*If the encrypted message appears then it originated with the Server*) |
293 |
lemma B_trusts_NS3: |
|
294 |
"\<lbrakk>Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs); |
|
11150 | 295 |
B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
296 |
\<Longrightarrow> \<exists>NA. Says Server A |
|
11104 | 297 |
(Crypt (shrK A) \<lbrace>NA, Agent B, Key K, |
298 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) |
|
299 |
\<in> set evs" |
|
300 |
apply (erule rev_mp) |
|
301 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies) |
|
302 |
apply auto |
|
303 |
done |
|
304 |
||
305 |
||
306 |
lemma A_trusts_NS4_lemma [rule_format]: |
|
307 |
"evs \<in> ns_shared \<Longrightarrow> |
|
308 |
Key K \<notin> analz (spies evs) \<longrightarrow> |
|
309 |
Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow> |
|
310 |
Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow> |
|
311 |
Says B A (Crypt K (Nonce NB)) \<in> set evs" |
|
312 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies) |
|
313 |
apply analz_mono_contra |
|
314 |
apply simp_all |
|
315 |
apply blast (*Fake*) |
|
316 |
(*NS2: contradiction from the assumptions |
|
317 |
Key K \<notin> used evs2 and Crypt K (Nonce NB) \<in> parts (spies evs2) *) |
|
318 |
apply (force dest!: Crypt_imp_keysFor) |
|
319 |
apply blast (*NS3*) |
|
320 |
(*NS4*) |
|
11150 | 321 |
apply (blast dest!: B_trusts_NS3 |
322 |
dest: Says_imp_knows_Spy [THEN analz.Inj] |
|
323 |
Crypt_Spy_analz_bad unique_session_keys) |
|
11104 | 324 |
done |
325 |
||
326 |
(*This version no longer assumes that K is secure*) |
|
327 |
lemma A_trusts_NS4: |
|
328 |
"\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs); |
|
329 |
Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); |
|
330 |
\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; |
|
331 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
332 |
\<Longrightarrow> Says B A (Crypt K (Nonce NB)) \<in> set evs" |
|
333 |
by (blast intro: A_trusts_NS4_lemma |
|
334 |
dest: A_trusts_NS2 Spy_not_see_encrypted_key) |
|
335 |
||
336 |
(*If the session key has been used in NS4 then somebody has forwarded |
|
337 |
component X in some instance of NS4. Perhaps an interesting property, |
|
338 |
but not needed (after all) for the proofs below.*) |
|
339 |
theorem NS4_implies_NS3 [rule_format]: |
|
340 |
"evs \<in> ns_shared \<Longrightarrow> |
|
341 |
Key K \<notin> analz (spies evs) \<longrightarrow> |
|
342 |
Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow> |
|
343 |
Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow> |
|
344 |
(\<exists>A'. Says A' B X \<in> set evs)" |
|
345 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies) |
|
346 |
apply analz_mono_contra |
|
347 |
apply (simp_all add: ex_disj_distrib) |
|
348 |
apply blast (*Fake*) |
|
349 |
apply (blast dest!: new_keys_not_used Crypt_imp_keysFor) (*NS2*) |
|
350 |
apply blast (*NS3*) |
|
351 |
(*NS4*) |
|
11150 | 352 |
apply (blast dest!: B_trusts_NS3 |
353 |
dest: Says_imp_knows_Spy [THEN analz.Inj] |
|
354 |
unique_session_keys Crypt_Spy_analz_bad) |
|
11104 | 355 |
done |
356 |
||
357 |
||
358 |
lemma B_trusts_NS5_lemma [rule_format]: |
|
359 |
"\<lbrakk>B \<notin> bad; evs \<in> ns_shared\<rbrakk> \<Longrightarrow> |
|
360 |
Key K \<notin> analz (spies evs) \<longrightarrow> |
|
361 |
Says Server A |
|
362 |
(Crypt (shrK A) \<lbrace>NA, Agent B, Key K, |
|
363 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow> |
|
364 |
Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow> |
|
365 |
Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs" |
|
366 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies) |
|
367 |
apply analz_mono_contra |
|
368 |
apply simp_all |
|
369 |
apply blast (*Fake*) |
|
370 |
apply (blast dest!: new_keys_not_used Crypt_imp_keysFor) (*NS2*) |
|
371 |
apply (blast dest!: cert_A_form) (*NS3*) |
|
372 |
(*NS5*) |
|
11150 | 373 |
apply (blast dest!: A_trusts_NS2 |
374 |
dest: Says_imp_knows_Spy [THEN analz.Inj] |
|
375 |
unique_session_keys Crypt_Spy_analz_bad) |
|
11104 | 376 |
done |
377 |
||
378 |
||
379 |
(*Very strong Oops condition reveals protocol's weakness*) |
|
380 |
lemma B_trusts_NS5: |
|
381 |
"\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs); |
|
382 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs); |
|
383 |
\<forall>NA NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; |
|
384 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
385 |
\<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs" |
|
11150 | 386 |
by (blast intro: B_trusts_NS5_lemma |
387 |
dest: B_trusts_NS3 Spy_not_see_encrypted_key) |
|
1934 | 388 |
|
389 |
end |