1 (* Title: HOL/Auth/NS_Shared.thy
2 Author: Lawrence C Paulson and Giampaolo Bella
3 Copyright 1996 University of Cambridge
6 header{*Needham-Schroeder Shared-Key Protocol and the Issues Property*}
8 theory NS_Shared imports Sledgehammer2d(*###*) Public begin
12 Burrows, Abadi and Needham (1989). A Logic of Authentication.
17 (* A is the true creator of X if she has sent X and X never appeared on
18 the trace before this event. Recall that traces grow from head. *)
19 Issues :: "[agent, agent, msg, event list] => bool"
20 ("_ Issues _ with _ on _") where
21 "A Issues B with X on evs =
22 (\<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
23 X \<notin> parts (spies (takeWhile (% z. z \<noteq> Says A B Y) (rev evs))))"
26 inductive_set ns_shared :: "event list set"
28 (*Initial trace is empty*)
29 Nil: "[] \<in> ns_shared"
30 (*The spy MAY say anything he CAN say. We do not expect him to
31 invent new nonces here, but he can also use NS1. Common to
32 all similar protocols.*)
33 | Fake: "\<lbrakk>evsf \<in> ns_shared; X \<in> synth (analz (spies evsf))\<rbrakk>
34 \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
36 (*Alice initiates a protocol run, requesting to talk to any B*)
37 | NS1: "\<lbrakk>evs1 \<in> ns_shared; Nonce NA \<notin> used evs1\<rbrakk>
38 \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared"
40 (*Server's response to Alice's message.
41 !! It may respond more than once to A's request !!
42 Server doesn't know who the true sender is, hence the A' in
44 | NS2: "\<lbrakk>evs2 \<in> ns_shared; Key KAB \<notin> used evs2; KAB \<in> symKeys;
45 Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
46 \<Longrightarrow> Says Server A
48 \<lbrace>Nonce NA, Agent B, Key KAB,
49 (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
50 # evs2 \<in> ns_shared"
52 (*We can't assume S=Server. Agent A "remembers" her nonce.
53 Need A \<noteq> Server because we allow messages to self.*)
54 | NS3: "\<lbrakk>evs3 \<in> ns_shared; A \<noteq> Server;
55 Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
56 Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
57 \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
59 (*Bob's nonce exchange. He does not know who the message came
60 from, but responds to A because she is mentioned inside.*)
61 | NS4: "\<lbrakk>evs4 \<in> ns_shared; Nonce NB \<notin> used evs4; K \<in> symKeys;
62 Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
63 \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
65 (*Alice responds with Nonce NB if she has seen the key before.
66 Maybe should somehow check Nonce NA again.
67 We do NOT send NB-1 or similar as the Spy cannot spoof such things.
68 Letting the Spy add or subtract 1 lets him send all nonces.
69 Instead we distinguish the messages by sending the nonce twice.*)
70 | NS5: "\<lbrakk>evs5 \<in> ns_shared; K \<in> symKeys;
71 Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
72 Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
73 \<in> set evs5\<rbrakk>
74 \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
76 (*This message models possible leaks of session keys.
77 The two Nonces identify the protocol run: the rule insists upon
78 the true senders in order to make them accurate.*)
79 | Oops: "\<lbrakk>evso \<in> ns_shared; Says B A (Crypt K (Nonce NB)) \<in> set evso;
80 Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
81 \<in> set evso\<rbrakk>
82 \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
85 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
86 declare parts.Body [dest]
87 declare Fake_parts_insert_in_Un [dest]
88 declare analz_into_parts [dest]
89 declare image_eq_UN [simp] (*accelerates proofs involving nested images*)
92 text{*A "possibility property": there are traces that reach the end*}
93 lemma "[| A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys |]
94 ==> \<exists>N. \<exists>evs \<in> ns_shared.
95 Says A B (Crypt K \<lbrace>Nonce N, Nonce N\<rbrace>) \<in> set evs"
96 apply (intro exI bexI)
97 apply (rule_tac [2] ns_shared.Nil
98 [THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3,
99 THEN ns_shared.NS4, THEN ns_shared.NS5])
100 apply (possibility, simp add: used_Cons)
103 (*This version is similar, while instantiating ?K and ?N to epsilon-terms
104 lemma "A \<noteq> Server \<Longrightarrow> \<exists>evs \<in> ns_shared.
105 Says A B (Crypt ?K \<lbrace>Nonce ?N, Nonce ?N\<rbrace>) \<in> set evs"
109 subsection{*Inductive proofs about @{term ns_shared}*}
111 subsubsection{*Forwarding lemmas, to aid simplification*}
113 text{*For reasoning about the encrypted portion of message NS3*}
114 lemma NS3_msg_in_parts_spies:
115 "Says S A (Crypt KA \<lbrace>N, B, K, X\<rbrace>) \<in> set evs \<Longrightarrow> X \<in> parts (spies evs)"
118 text{*For reasoning about the Oops message*}
119 lemma Oops_parts_spies:
120 "Says Server A (Crypt (shrK A) \<lbrace>NA, B, K, X\<rbrace>) \<in> set evs
121 \<Longrightarrow> K \<in> parts (spies evs)"
124 text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that NOBODY
125 sends messages containing @{term X}*}
127 text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
128 lemma Spy_see_shrK [simp]:
129 "evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
130 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all, blast+)
133 lemma Spy_analz_shrK [simp]:
134 "evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
138 text{*Nobody can have used non-existent keys!*}
139 lemma new_keys_not_used [simp]:
140 "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> ns_shared|]
141 ==> K \<notin> keysFor (parts (spies evs))"
143 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all)
144 txt{*Fake, NS2, NS4, NS5*}
145 apply (force dest!: keysFor_parts_insert, blast+)
149 subsubsection{*Lemmas concerning the form of items passed in messages*}
151 text{*Describes the form of K, X and K' when the Server sends this message.*}
152 lemma Says_Server_message_form:
153 "\<lbrakk>Says Server A (Crypt K' \<lbrace>N, Agent B, Key K, X\<rbrace>) \<in> set evs;
154 evs \<in> ns_shared\<rbrakk>
155 \<Longrightarrow> K \<notin> range shrK \<and>
156 X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<and>
158 by (erule rev_mp, erule ns_shared.induct, auto)
161 text{*If the encrypted message appears then it originated with the Server*}
163 "\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
164 A \<notin> bad; evs \<in> ns_shared\<rbrakk>
165 \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs"
167 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, auto)
171 "\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
172 A \<notin> bad; evs \<in> ns_shared\<rbrakk>
173 \<Longrightarrow> K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)"
174 by (blast dest!: A_trusts_NS2 Says_Server_message_form)
176 text{*EITHER describes the form of X when the following message is sent,
177 OR reduces it to the Fake case.
178 Use @{text Says_Server_message_form} if applicable.*}
179 lemma Says_S_message_form:
180 "\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
181 evs \<in> ns_shared\<rbrakk>
182 \<Longrightarrow> (K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>))
183 \<or> X \<in> analz (spies evs)"
184 by (blast dest: Says_imp_knows_Spy analz_shrK_Decrypt cert_A_form analz.Inj)
187 (*Alternative version also provable
188 lemma Says_S_message_form2:
189 "\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
190 evs \<in> ns_shared\<rbrakk>
191 \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs
192 \<or> X \<in> analz (spies evs)"
193 apply (case_tac "A \<in> bad")
194 apply (force dest!: Says_imp_knows_Spy [THEN analz.Inj])
195 by (blast dest!: A_trusts_NS2 Says_Server_message_form)
200 SESSION KEY COMPROMISE THEOREM. To prove theorems of the form
202 Key K \<in> analz (insert (Key KAB) (spies evs)) \<Longrightarrow>
203 Key K \<in> analz (spies evs)
205 A more general formula must be proved inductively.
208 text{*NOT useful in this form, but it says that session keys are not used
209 to encrypt messages containing other keys, in the actual protocol.
210 We require that agents should behave like this subsequently also.*}
211 lemma "\<lbrakk>evs \<in> ns_shared; Kab \<notin> range shrK\<rbrakk> \<Longrightarrow>
212 (Crypt KAB X) \<in> parts (spies evs) \<and>
213 Key K \<in> parts {X} \<longrightarrow> Key K \<in> parts (spies evs)"
214 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all)
216 apply (blast dest: parts_insert_subset_Un)
217 txt{*Base, NS4 and NS5*}
222 subsubsection{*Session keys are not used to encrypt other session keys*}
224 text{*The equality makes the induction hypothesis easier to apply*}
226 lemma analz_image_freshK [rule_format]:
227 "evs \<in> ns_shared \<Longrightarrow>
228 \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
229 (Key K \<in> analz (Key`KK \<union> (spies evs))) =
230 (K \<in> KK \<or> Key K \<in> analz (spies evs))"
231 apply (erule ns_shared.induct)
232 apply (drule_tac [8] Says_Server_message_form)
233 apply (erule_tac [5] Says_S_message_form [THEN disjE], analz_freshK, spy_analz)
239 lemma analz_insert_freshK:
240 "\<lbrakk>evs \<in> ns_shared; KAB \<notin> range shrK\<rbrakk> \<Longrightarrow>
241 (Key K \<in> analz (insert (Key KAB) (spies evs))) =
242 (K = KAB \<or> Key K \<in> analz (spies evs))"
243 by (simp only: analz_image_freshK analz_image_freshK_simps)
246 subsubsection{*The session key K uniquely identifies the message*}
248 text{*In messages of this form, the session key uniquely identifies the rest*}
249 lemma unique_session_keys:
250 "\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
251 Says Server A' (Crypt (shrK A') \<lbrace>NA', Agent B', Key K, X'\<rbrace>) \<in> set evs;
252 evs \<in> ns_shared\<rbrakk> \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B' \<and> X = X'"
253 by (erule rev_mp, erule rev_mp, erule ns_shared.induct, simp_all, blast+)
256 subsubsection{*Crucial secrecy property: Spy doesn't see the keys sent in NS2*}
258 text{*Beware of @{text "[rule_format]"} and the universal quantifier!*}
260 "\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
261 Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>)
263 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk>
264 \<Longrightarrow> (\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs) \<longrightarrow>
265 Key K \<notin> analz (spies evs)"
267 apply (erule ns_shared.induct, force)
268 apply (frule_tac [7] Says_Server_message_form)
269 apply (frule_tac [4] Says_S_message_form)
270 apply (erule_tac [5] disjE)
271 apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs, spy_analz)
275 apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2
276 dest: Says_imp_knows_Spy analz.Inj unique_session_keys)
278 apply (blast dest: unique_session_keys)
283 text{*Final version: Server's message in the most abstract form*}
284 lemma Spy_not_see_encrypted_key:
285 "\<lbrakk>Says Server A (Crypt K' \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs;
286 \<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
287 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk>
288 \<Longrightarrow> Key K \<notin> analz (spies evs)"
289 by (blast dest: Says_Server_message_form secrecy_lemma)
292 subsection{*Guarantees available at various stages of protocol*}
294 text{*If the encrypted message appears then it originated with the Server*}
296 "\<lbrakk>Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs);
297 B \<notin> bad; evs \<in> ns_shared\<rbrakk>
298 \<Longrightarrow> \<exists>NA. Says Server A
299 (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
300 Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>)
303 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, auto)
307 lemma A_trusts_NS4_lemma [rule_format]:
308 "evs \<in> ns_shared \<Longrightarrow>
309 Key K \<notin> analz (spies evs) \<longrightarrow>
310 Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow>
311 Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
312 Says B A (Crypt K (Nonce NB)) \<in> set evs"
313 apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies)
314 sledgehammer [atp = e, overlord] (Crypt_synth Fake_parts_insert_in_Un List.set.simps(2) Un_commute agent.simps(4) analz_insertI event.simps(1) insertE insert_iff knows_Spy_Says mem_def not_parts_not_analz parts_analz parts_synth sup1E)
315 apply (metis Crypt_synth Fake_parts_insert_in_Un List.set.simps(2) Un_commute agent.simps(4) analz_insertI event.simps(1) insertE insert_iff knows_Spy_Says mem_def not_parts_not_analz parts_analz parts_synth sup1E)
316 apply (analz_mono_contra, simp_all, blast)
317 txt{*NS2: contradiction from the assumptions @{term "Key K \<notin> used evs2"} and
318 @{term "Crypt K (Nonce NB) \<in> parts (spies evs2)"} *}
319 apply (force dest!: Crypt_imp_keysFor)
321 apply (metis B_trusts_NS3 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst unique_session_keys)
324 text{*This version no longer assumes that K is secure*}
326 "\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs);
327 Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
328 \<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
329 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk>
330 \<Longrightarrow> Says B A (Crypt K (Nonce NB)) \<in> set evs"
331 by (blast intro: A_trusts_NS4_lemma
332 dest: A_trusts_NS2 Spy_not_see_encrypted_key)
334 text{*If the session key has been used in NS4 then somebody has forwarded
335 component X in some instance of NS4. Perhaps an interesting property,
336 but not needed (after all) for the proofs below.*}
337 theorem NS4_implies_NS3 [rule_format]:
338 "evs \<in> ns_shared \<Longrightarrow>
339 Key K \<notin> analz (spies evs) \<longrightarrow>
340 Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow>
341 Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
342 (\<exists>A'. Says A' B X \<in> set evs)"
343 apply (erule ns_shared.induct, force)
344 apply (drule_tac [4] NS3_msg_in_parts_spies)
345 apply analz_mono_contra
346 apply (simp_all add: ex_disj_distrib, blast)
348 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
350 apply (metis B_trusts_NS3 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst unique_session_keys)
354 lemma B_trusts_NS5_lemma [rule_format]:
355 "\<lbrakk>B \<notin> bad; evs \<in> ns_shared\<rbrakk> \<Longrightarrow>
356 Key K \<notin> analz (spies evs) \<longrightarrow>
358 (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
359 Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
360 Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
361 Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"
362 apply (erule ns_shared.induct, force)
363 apply (drule_tac [4] NS3_msg_in_parts_spies)
364 apply (analz_mono_contra, simp_all, blast)
366 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
368 apply (blast dest!: A_trusts_NS2
369 dest: Says_imp_knows_Spy [THEN analz.Inj]
370 unique_session_keys Crypt_Spy_analz_bad)
374 text{*Very strong Oops condition reveals protocol's weakness*}
376 "\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs);
377 Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs);
378 \<forall>NA NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs;
379 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk>
380 \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"
381 by (blast intro: B_trusts_NS5_lemma
382 dest: B_trusts_NS3 Spy_not_see_encrypted_key)
384 text{*Unaltered so far wrt original version*}
386 subsection{*Lemmas for reasoning about predicate "Issues"*}
388 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
389 apply (induct_tac "evs")
390 apply (induct_tac [2] "a", auto)
393 lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
394 apply (induct_tac "evs")
395 apply (induct_tac [2] "a", auto)
398 lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
399 (if A:bad then insert X (spies evs) else spies evs)"
400 apply (induct_tac "evs")
401 apply (induct_tac [2] "a", auto)
404 lemma spies_evs_rev: "spies evs = spies (rev evs)"
405 apply (induct_tac "evs")
406 apply (induct_tac [2] "a")
407 apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
410 lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono]
412 lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs"
413 apply (induct_tac "evs")
414 apply (induct_tac [2] "a", auto)
415 txt{* Resembles @{text"used_subset_append"} in theory Event.*}
418 lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
421 subsection{*Guarantees of non-injective agreement on the session key, and
422 of key distribution. They also express forms of freshness of certain messages,
423 namely that agents were alive after something happened.*}
426 "\<lbrakk> Says B A (Crypt K (Nonce Nb)) \<in> set evs;
427 Key K \<notin> analz (spies evs);
428 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared \<rbrakk>
429 \<Longrightarrow> B Issues A with (Crypt K (Nonce Nb)) on evs"
430 apply (simp (no_asm) add: Issues_def)
432 apply (rule conjI, assumption)
433 apply (simp (no_asm))
436 apply (erule ns_shared.induct, analz_mono_contra)
440 apply (simp_all add: takeWhile_tail)
441 txt{*NS3 remains by pure coincidence!*}
442 apply (force dest!: A_trusts_NS2 Says_Server_message_form)
443 txt{*NS4 would be the non-trivial case can be solved by Nb being used*}
444 apply (blast dest: parts_spies_takeWhile_mono [THEN subsetD]
445 parts_spies_evs_revD2 [THEN subsetD])
448 text{*Tells A that B was alive after she sent him the session key. The
449 session key must be assumed confidential for this deduction to be meaningful,
450 but that assumption can be relaxed by the appropriate argument.
452 Precisely, the theorem guarantees (to A) key distribution of the session key
453 to B. It also guarantees (to A) non-injective agreement of B with A on the
454 session key. Both goals are available to A in the sense of Goal Availability.
456 lemma A_authenticates_and_keydist_to_B:
457 "\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs);
458 Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
459 Key K \<notin> analz(knows Spy evs);
460 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk>
461 \<Longrightarrow> B Issues A with (Crypt K (Nonce NB)) on evs"
462 by (blast intro: A_trusts_NS4_lemma B_Issues_A dest: A_trusts_NS2)
465 "\<lbrakk> Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts(spies evs);
466 Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace> \<in> parts(spies evs);
467 Key K \<notin> analz (spies evs);
468 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared \<rbrakk>
469 \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs";
473 apply (erule ns_shared.induct, analz_mono_contra)
478 apply (force dest!: Crypt_imp_keysFor)
480 apply (metis NS3_msg_in_parts_spies parts_cut_eq)
481 txt{*NS5, the most important case, can be solved by unicity*}
482 apply (metis A_trusts_NS2 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst analz.Snd unique_session_keys)
486 "\<lbrakk> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs;
487 Key K \<notin> analz (spies evs);
488 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared \<rbrakk>
489 \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) on evs"
490 apply (simp (no_asm) add: Issues_def)
492 apply (rule conjI, assumption)
493 apply (simp (no_asm))
496 apply (erule ns_shared.induct, analz_mono_contra)
500 apply (simp_all add: takeWhile_tail)
501 txt{*NS3 remains by pure coincidence!*}
502 apply (force dest!: A_trusts_NS2 Says_Server_message_form)
503 txt{*NS5 is the non-trivial case and cannot be solved as in @{term B_Issues_A}! because NB is not fresh. We need @{term A_trusts_NS5}, proved for this very purpose*}
504 apply (blast dest: A_trusts_NS5 parts_spies_takeWhile_mono [THEN subsetD]
505 parts_spies_evs_revD2 [THEN subsetD])
508 text{*Tells B that A was alive after B issued NB.
510 Precisely, the theorem guarantees (to B) key distribution of the session key to A. It also guarantees (to B) non-injective agreement of A with B on the session key. Both goals are available to B in the sense of Goal Availability.
512 lemma B_authenticates_and_keydist_to_A:
513 "\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs);
514 Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs);
515 Key K \<notin> analz (spies evs);
516 A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk>
517 \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) on evs"
518 by (blast intro: A_Issues_B B_trusts_NS5_lemma dest: B_trusts_NS3)