author | paulson |
Tue, 23 Sep 2003 15:41:33 +0200 | |
changeset 14200 | d8598e24f8fa |
parent 13956 | 8fe7e12290e1 |
child 14207 | f20fbb141673 |
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*) |
|
13926 | 19 |
Nil: "[] \<in> ns_shared" |
11104 | 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.*) |
|
13926 | 23 |
Fake: "\<lbrakk>evsf \<in> ns_shared; X \<in> synth (analz (spies evsf))\<rbrakk> |
24 |
\<Longrightarrow> Says Spy B X # evsf \<in> ns_shared" |
|
11104 | 25 |
|
26 |
(*Alice initiates a protocol run, requesting to talk to any B*) |
|
13926 | 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" |
|
11104 | 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.*) |
|
13926 | 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 |
|
11104 | 37 |
(Crypt (shrK A) |
13926 | 38 |
\<lbrace>Nonce NA, Agent B, Key KAB, |
39 |
(Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>) |
|
40 |
# evs2 \<in> ns_shared" |
|
11104 | 41 |
|
42 |
(*We can't assume S=Server. Agent A "remembers" her nonce. |
|
13926 | 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" |
|
11104 | 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.*) |
|
13926 | 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. |
|
11465 | 58 |
Letting the Spy add or subtract 1 lets him send all nonces. |
11104 | 59 |
Instead we distinguish the messages by sending the nonce twice.*) |
13926 | 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" |
|
11104 | 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.*) |
|
13926 | 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" |
|
11104 | 73 |
|
11150 | 74 |
|
75 |
declare Says_imp_knows_Spy [THEN parts.Inj, dest] |
|
76 |
declare parts.Body [dest] |
|
11251 | 77 |
declare Fake_parts_insert_in_Un [dest] |
78 |
declare analz_into_parts [dest] |
|
11104 | 79 |
declare image_eq_UN [simp] (*accelerates proofs involving nested images*) |
80 |
||
81 |
||
13926 | 82 |
text{*A "possibility property": there are traces that reach the end*} |
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
83 |
lemma "[| A \<noteq> Server; Key K \<notin> used [] |] |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
84 |
==> \<exists>N. \<exists>evs \<in> ns_shared. |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
85 |
Says A B (Crypt K \<lbrace>Nonce N, Nonce N\<rbrace>) \<in> set evs" |
11104 | 86 |
apply (intro exI bexI) |
87 |
apply (rule_tac [2] ns_shared.Nil |
|
88 |
[THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3, |
|
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
89 |
THEN ns_shared.NS4, THEN ns_shared.NS5]) |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
90 |
apply (possibility, simp add: used_Cons) |
11104 | 91 |
done |
92 |
||
93 |
(*This version is similar, while instantiating ?K and ?N to epsilon-terms |
|
13926 | 94 |
lemma "A \<noteq> Server \<Longrightarrow> \<exists>evs \<in> ns_shared. |
95 |
Says A B (Crypt ?K \<lbrace>Nonce ?N, Nonce ?N\<rbrace>) \<in> set evs" |
|
11104 | 96 |
*) |
97 |
||
98 |
||
13926 | 99 |
subsection{*Inductive proofs about @{term ns_shared}*} |
11104 | 100 |
|
13926 | 101 |
subsubsection{*Forwarding lemmas, to aid simplification*} |
1934 | 102 |
|
13926 | 103 |
text{*For reasoning about the encrypted portion of message NS3*} |
11104 | 104 |
lemma NS3_msg_in_parts_spies: |
13926 | 105 |
"Says S A (Crypt KA \<lbrace>N, B, K, X\<rbrace>) \<in> set evs \<Longrightarrow> X \<in> parts (spies evs)" |
11104 | 106 |
by blast |
11280 | 107 |
|
13926 | 108 |
text{*For reasoning about the Oops message*} |
11104 | 109 |
lemma Oops_parts_spies: |
13926 | 110 |
"Says Server A (Crypt (shrK A) \<lbrace>NA, B, K, X\<rbrace>) \<in> set evs |
111 |
\<Longrightarrow> K \<in> parts (spies evs)" |
|
11104 | 112 |
by blast |
113 |
||
13926 | 114 |
text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that NOBODY |
115 |
sends messages containing @{term X}*} |
|
11104 | 116 |
|
13926 | 117 |
text{*Spy never sees another agent's shared key! (unless it's bad at start)*} |
11104 | 118 |
lemma Spy_see_shrK [simp]: |
13926 | 119 |
"evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" |
13507 | 120 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all, blast+) |
11104 | 121 |
done |
122 |
||
123 |
lemma Spy_analz_shrK [simp]: |
|
13926 | 124 |
"evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)" |
11104 | 125 |
by auto |
126 |
||
127 |
||
13926 | 128 |
text{*Nobody can have used non-existent keys!*} |
11104 | 129 |
lemma new_keys_not_used [rule_format, simp]: |
13926 | 130 |
"evs \<in> ns_shared \<Longrightarrow> Key K \<notin> used evs \<longrightarrow> K \<notin> keysFor (parts (spies evs))" |
13507 | 131 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all) |
13926 | 132 |
txt{*Fake, NS2, NS4, NS5*} |
133 |
apply (force dest!: keysFor_parts_insert, blast+) |
|
11104 | 134 |
done |
135 |
||
136 |
||
13926 | 137 |
subsubsection{*Lemmas concerning the form of items passed in messages*} |
11104 | 138 |
|
13926 | 139 |
text{*Describes the form of K, X and K' when the Server sends this message.*} |
11104 | 140 |
lemma Says_Server_message_form: |
13926 | 141 |
"\<lbrakk>Says Server A (Crypt K' \<lbrace>N, Agent B, Key K, X\<rbrace>) \<in> set evs; |
142 |
evs \<in> ns_shared\<rbrakk> |
|
143 |
\<Longrightarrow> K \<notin> range shrK \<and> |
|
144 |
X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<and> |
|
11104 | 145 |
K' = shrK A" |
146 |
by (erule rev_mp, erule ns_shared.induct, auto) |
|
147 |
||
1934 | 148 |
|
13926 | 149 |
text{*If the encrypted message appears then it originated with the Server*} |
11104 | 150 |
lemma A_trusts_NS2: |
13926 | 151 |
"\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); |
152 |
A \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
153 |
\<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs" |
|
11104 | 154 |
apply (erule rev_mp) |
13507 | 155 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, auto) |
11104 | 156 |
done |
157 |
||
158 |
lemma cert_A_form: |
|
13926 | 159 |
"\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); |
160 |
A \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
161 |
\<Longrightarrow> K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)" |
|
11104 | 162 |
by (blast dest!: A_trusts_NS2 Says_Server_message_form) |
163 |
||
164 |
(*EITHER describes the form of X when the following message is sent, |
|
165 |
OR reduces it to the Fake case. |
|
166 |
Use Says_Server_message_form if applicable.*) |
|
167 |
lemma Says_S_message_form: |
|
13926 | 168 |
"\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
169 |
evs \<in> ns_shared\<rbrakk> |
|
170 |
\<Longrightarrow> (K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)) |
|
171 |
\<or> X \<in> analz (spies evs)" |
|
11150 | 172 |
by (blast dest: Says_imp_knows_Spy cert_A_form analz.Inj) |
173 |
||
11104 | 174 |
|
175 |
(*Alternative version also provable |
|
176 |
lemma Says_S_message_form2: |
|
13926 | 177 |
"\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
178 |
evs \<in> ns_shared\<rbrakk> |
|
179 |
\<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs |
|
180 |
\<or> X \<in> analz (spies evs)" |
|
181 |
apply (case_tac "A \<in> bad") |
|
13507 | 182 |
apply (force dest!: Says_imp_knows_Spy [THEN analz.Inj]) |
11104 | 183 |
by (blast dest!: A_trusts_NS2 Says_Server_message_form) |
184 |
*) |
|
185 |
||
186 |
||
187 |
(**** |
|
188 |
SESSION KEY COMPROMISE THEOREM. To prove theorems of the form |
|
189 |
||
13926 | 190 |
Key K \<in> analz (insert (Key KAB) (spies evs)) \<Longrightarrow> |
191 |
Key K \<in> analz (spies evs) |
|
11104 | 192 |
|
193 |
A more general formula must be proved inductively. |
|
194 |
****) |
|
1934 | 195 |
|
13926 | 196 |
text{*NOT useful in this form, but it says that session keys are not used |
11104 | 197 |
to encrypt messages containing other keys, in the actual protocol. |
13926 | 198 |
We require that agents should behave like this subsequently also.*} |
199 |
lemma "\<lbrakk>evs \<in> ns_shared; Kab \<notin> range shrK\<rbrakk> \<Longrightarrow> |
|
200 |
(Crypt KAB X) \<in> parts (spies evs) \<and> |
|
201 |
Key K \<in> parts {X} \<longrightarrow> Key K \<in> parts (spies evs)" |
|
13507 | 202 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all) |
13926 | 203 |
txt{*Fake*} |
11104 | 204 |
apply (blast dest: parts_insert_subset_Un) |
13926 | 205 |
txt{*Base, NS4 and NS5*} |
11104 | 206 |
apply auto |
207 |
done |
|
208 |
||
209 |
||
13926 | 210 |
subsubsection{*Session keys are not used to encrypt other session keys*} |
11104 | 211 |
|
13926 | 212 |
text{*The equality makes the induction hypothesis easier to apply*} |
11104 | 213 |
|
214 |
lemma analz_image_freshK [rule_format]: |
|
13926 | 215 |
"evs \<in> ns_shared \<Longrightarrow> |
216 |
\<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow> |
|
217 |
(Key K \<in> analz (Key`KK \<union> (spies evs))) = |
|
218 |
(K \<in> KK \<or> Key K \<in> analz (spies evs))" |
|
11104 | 219 |
apply (erule ns_shared.induct, force) |
11280 | 220 |
apply (drule_tac [7] Says_Server_message_form) |
13507 | 221 |
apply (erule_tac [4] Says_S_message_form [THEN disjE], analz_freshK, spy_analz) |
11104 | 222 |
done |
223 |
||
224 |
||
225 |
lemma analz_insert_freshK: |
|
13926 | 226 |
"\<lbrakk>evs \<in> ns_shared; KAB \<notin> range shrK\<rbrakk> \<Longrightarrow> |
227 |
(Key K \<in> analz (insert (Key KAB) (spies evs))) = |
|
228 |
(K = KAB \<or> Key K \<in> analz (spies evs))" |
|
11104 | 229 |
by (simp only: analz_image_freshK analz_image_freshK_simps) |
230 |
||
231 |
||
13926 | 232 |
subsubsection{*The session key K uniquely identifies the message*} |
1934 | 233 |
|
13926 | 234 |
text{*In messages of this form, the session key uniquely identifies the rest*} |
11104 | 235 |
lemma unique_session_keys: |
13926 | 236 |
"\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
237 |
Says Server A' (Crypt (shrK A') \<lbrace>NA', Agent B', Key K, X'\<rbrace>) \<in> set evs; |
|
238 |
evs \<in> ns_shared\<rbrakk> \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B' \<and> X = X'" |
|
13507 | 239 |
apply (erule rev_mp, erule rev_mp, erule ns_shared.induct, simp_all, blast+) |
11104 | 240 |
done |
241 |
||
242 |
||
13926 | 243 |
subsubsection{*Crucial secrecy property: Spy does not see the keys sent in msg NS2*} |
11104 | 244 |
|
13956 | 245 |
text{*Beware of @{text "[rule_format]"} and the universal quantifier!*} |
11150 | 246 |
lemma secrecy_lemma: |
13926 | 247 |
"\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, |
248 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) |
|
249 |
\<in> set evs; |
|
250 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
251 |
\<Longrightarrow> (\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs) \<longrightarrow> |
|
252 |
Key K \<notin> analz (spies evs)" |
|
11104 | 253 |
apply (erule rev_mp) |
254 |
apply (erule ns_shared.induct, force) |
|
255 |
apply (frule_tac [7] Says_Server_message_form) |
|
256 |
apply (frule_tac [4] Says_S_message_form) |
|
257 |
apply (erule_tac [5] disjE) |
|
13926 | 258 |
apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs, spy_analz) |
259 |
txt{*NS2*} |
|
260 |
apply blast |
|
261 |
txt{*NS3, Server sub-case*} |
|
11188 | 262 |
apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2 |
263 |
dest: Says_imp_knows_Spy analz.Inj unique_session_keys) |
|
13926 | 264 |
txt{*NS3, Spy sub-case; also Oops*} |
11280 | 265 |
apply (blast dest: unique_session_keys)+ |
11104 | 266 |
done |
267 |
||
268 |
||
11188 | 269 |
|
13926 | 270 |
text{*Final version: Server's message in the most abstract form*} |
11104 | 271 |
lemma Spy_not_see_encrypted_key: |
13926 | 272 |
"\<lbrakk>Says Server A (Crypt K' \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs; |
273 |
\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; |
|
274 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
275 |
\<Longrightarrow> Key K \<notin> analz (spies evs)" |
|
11150 | 276 |
by (blast dest: Says_Server_message_form secrecy_lemma) |
11104 | 277 |
|
278 |
||
13926 | 279 |
subsection{*Guarantees available at various stages of protocol*} |
1934 | 280 |
|
13926 | 281 |
text{*If the encrypted message appears then it originated with the Server*} |
11104 | 282 |
lemma B_trusts_NS3: |
13926 | 283 |
"\<lbrakk>Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs); |
284 |
B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
285 |
\<Longrightarrow> \<exists>NA. Says Server A |
|
286 |
(Crypt (shrK A) \<lbrace>NA, Agent B, Key K, |
|
287 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) |
|
288 |
\<in> set evs" |
|
11104 | 289 |
apply (erule rev_mp) |
13507 | 290 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, auto) |
11104 | 291 |
done |
292 |
||
293 |
||
294 |
lemma A_trusts_NS4_lemma [rule_format]: |
|
13926 | 295 |
"evs \<in> ns_shared \<Longrightarrow> |
296 |
Key K \<notin> analz (spies evs) \<longrightarrow> |
|
297 |
Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow> |
|
298 |
Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow> |
|
299 |
Says B A (Crypt K (Nonce NB)) \<in> set evs" |
|
11104 | 300 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies) |
13926 | 301 |
apply (analz_mono_contra, simp_all, blast) |
11280 | 302 |
(*NS2: contradiction from the assumptions |
13926 | 303 |
Key K \<notin> used evs2 and Crypt K (Nonce NB) \<in> parts (spies evs2) *) |
304 |
apply (force dest!: Crypt_imp_keysFor) |
|
305 |
txt{*NS3*} |
|
306 |
apply blast |
|
13935 | 307 |
txt{*NS4*} |
308 |
apply (blast dest: B_trusts_NS3 |
|
309 |
Says_imp_knows_Spy [THEN analz.Inj] |
|
11150 | 310 |
Crypt_Spy_analz_bad unique_session_keys) |
11104 | 311 |
done |
312 |
||
13926 | 313 |
text{*This version no longer assumes that K is secure*} |
11104 | 314 |
lemma A_trusts_NS4: |
13926 | 315 |
"\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs); |
316 |
Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); |
|
317 |
\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; |
|
318 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
319 |
\<Longrightarrow> Says B A (Crypt K (Nonce NB)) \<in> set evs" |
|
11280 | 320 |
by (blast intro: A_trusts_NS4_lemma |
11104 | 321 |
dest: A_trusts_NS2 Spy_not_see_encrypted_key) |
322 |
||
323 |
(*If the session key has been used in NS4 then somebody has forwarded |
|
11280 | 324 |
component X in some instance of NS4. Perhaps an interesting property, |
11104 | 325 |
but not needed (after all) for the proofs below.*) |
326 |
theorem NS4_implies_NS3 [rule_format]: |
|
13926 | 327 |
"evs \<in> ns_shared \<Longrightarrow> |
328 |
Key K \<notin> analz (spies evs) \<longrightarrow> |
|
329 |
Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow> |
|
330 |
Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow> |
|
331 |
(\<exists>A'. Says A' B X \<in> set evs)" |
|
13507 | 332 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, analz_mono_contra) |
13926 | 333 |
apply (simp_all add: ex_disj_distrib, blast) |
334 |
txt{*NS2*} |
|
335 |
apply (blast dest!: new_keys_not_used Crypt_imp_keysFor) |
|
336 |
txt{*NS3*} |
|
337 |
apply blast |
|
338 |
txt{*NS4*} |
|
13935 | 339 |
apply (blast dest: B_trusts_NS3 |
11280 | 340 |
dest: Says_imp_knows_Spy [THEN analz.Inj] |
11150 | 341 |
unique_session_keys Crypt_Spy_analz_bad) |
11104 | 342 |
done |
343 |
||
344 |
||
345 |
lemma B_trusts_NS5_lemma [rule_format]: |
|
13926 | 346 |
"\<lbrakk>B \<notin> bad; evs \<in> ns_shared\<rbrakk> \<Longrightarrow> |
347 |
Key K \<notin> analz (spies evs) \<longrightarrow> |
|
11104 | 348 |
Says Server A |
13926 | 349 |
(Crypt (shrK A) \<lbrace>NA, Agent B, Key K, |
350 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow> |
|
351 |
Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow> |
|
352 |
Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs" |
|
353 |
apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, analz_mono_contra, simp_all, blast) |
|
354 |
txt{*NS2*} |
|
355 |
apply (blast dest!: new_keys_not_used Crypt_imp_keysFor) |
|
356 |
txt{*NS3*} |
|
357 |
apply (blast dest!: cert_A_form) |
|
358 |
txt{*NS5*} |
|
11150 | 359 |
apply (blast dest!: A_trusts_NS2 |
11280 | 360 |
dest: Says_imp_knows_Spy [THEN analz.Inj] |
11150 | 361 |
unique_session_keys Crypt_Spy_analz_bad) |
11104 | 362 |
done |
363 |
||
364 |
||
13926 | 365 |
text{*Very strong Oops condition reveals protocol's weakness*} |
11104 | 366 |
lemma B_trusts_NS5: |
13926 | 367 |
"\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs); |
368 |
Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs); |
|
369 |
\<forall>NA NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; |
|
370 |
A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> |
|
371 |
\<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs" |
|
11280 | 372 |
by (blast intro: B_trusts_NS5_lemma |
11150 | 373 |
dest: B_trusts_NS3 Spy_not_see_encrypted_key) |
1934 | 374 |
|
375 |
end |