author | paulson |
Fri, 29 Nov 1996 18:03:21 +0100 | |
changeset 2284 | 80ebd1a213fd |
parent 2264 | f298678bd54a |
child 2375 | 14539397fc04 |
permissions | -rw-r--r-- |
2002 | 1 |
(* Title: HOL/Auth/OtwayRees_Bad |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
Inductive relation "otway" for the Otway-Rees protocol. |
|
7 |
||
8 |
The FAULTY version omitting encryption of Nonce NB, as suggested on page 247 of |
|
9 |
Burrows, Abadi and Needham. A Logic of Authentication. |
|
10 |
Proc. Royal Soc. 426 (1989) |
|
11 |
||
12 |
This file illustrates the consequences of such errors. We can still prove |
|
2032 | 13 |
impressive-looking properties such as Spy_not_see_encrypted_key, yet the |
2002 | 14 |
protocol is open to a middleperson attack. Attempting to prove some key lemmas |
15 |
indicates the possibility of this attack. |
|
16 |
*) |
|
17 |
||
18 |
open OtwayRees_Bad; |
|
19 |
||
20 |
proof_timing:=true; |
|
21 |
HOL_quantifiers := false; |
|
22 |
||
23 |
||
24 |
(*Weak liveness: there are traces that reach the end*) |
|
25 |
goal thy |
|
26 |
"!!A B. [| A ~= B; A ~= Server; B ~= Server |] \ |
|
27 |
\ ==> EX K. EX NA. EX evs: otway. \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
28 |
\ Says B A {|Nonce NA, Crypt (shrK A) {|Nonce NA, Key K|}|} \ |
2002 | 29 |
\ : set_of_list evs"; |
30 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
|
2032 | 31 |
by (rtac (otway.Nil RS otway.OR1 RS otway.OR2 RS otway.OR3 RS otway.OR4) 2); |
2002 | 32 |
by (ALLGOALS (simp_tac (!simpset setsolver safe_solver))); |
33 |
by (REPEAT_FIRST (resolve_tac [refl, conjI])); |
|
34 |
by (ALLGOALS (fast_tac (!claset addss (!simpset setsolver safe_solver)))); |
|
35 |
result(); |
|
36 |
||
37 |
||
38 |
(**** Inductive proofs about otway ****) |
|
39 |
||
2032 | 40 |
(*The Spy can see more than anybody else, except for their initial state*) |
2002 | 41 |
goal thy |
42 |
"!!evs. evs : otway ==> \ |
|
2052 | 43 |
\ sees lost A evs <= initState lost A Un sees lost Spy evs"; |
2032 | 44 |
by (etac otway.induct 1); |
2002 | 45 |
by (ALLGOALS (fast_tac (!claset addDs [sees_Says_subset_insert RS subsetD] |
2032 | 46 |
addss (!simpset)))); |
47 |
qed "sees_agent_subset_sees_Spy"; |
|
2002 | 48 |
|
49 |
||
50 |
(*Nobody sends themselves messages*) |
|
51 |
goal thy "!!evs. evs : otway ==> ALL A X. Says A A X ~: set_of_list evs"; |
|
2032 | 52 |
by (etac otway.induct 1); |
2002 | 53 |
by (Auto_tac()); |
54 |
qed_spec_mp "not_Says_to_self"; |
|
55 |
Addsimps [not_Says_to_self]; |
|
56 |
AddSEs [not_Says_to_self RSN (2, rev_notE)]; |
|
57 |
||
58 |
||
59 |
(** For reasoning about the encrypted portion of messages **) |
|
60 |
||
61 |
goal thy "!!evs. Says A' B {|N, Agent A, Agent B, X|} : set_of_list evs ==> \ |
|
2052 | 62 |
\ X : analz (sees lost Spy evs)"; |
2032 | 63 |
by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]) 1); |
64 |
qed "OR2_analz_sees_Spy"; |
|
2002 | 65 |
|
66 |
goal thy "!!evs. Says S B {|N, X, X'|} : set_of_list evs ==> \ |
|
2052 | 67 |
\ X : analz (sees lost Spy evs)"; |
2032 | 68 |
by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]) 1); |
69 |
qed "OR4_analz_sees_Spy"; |
|
2002 | 70 |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
71 |
goal thy "!!evs. Says Server B {|NA, X, Crypt K' {|NB,K|}|} : set_of_list evs \ |
2131 | 72 |
\ ==> K : parts (sees lost Spy evs)"; |
2002 | 73 |
by (fast_tac (!claset addSEs partsEs |
2032 | 74 |
addSDs [Says_imp_sees_Spy RS parts.Inj]) 1); |
2131 | 75 |
qed "Oops_parts_sees_Spy"; |
2002 | 76 |
|
77 |
(*OR2_analz... and OR4_analz... let us treat those cases using the same |
|
78 |
argument as for the Fake case. This is possible for most, but not all, |
|
79 |
proofs: Fake does not invent new nonces (as in OR2), and of course Fake |
|
2032 | 80 |
messages originate from the Spy. *) |
2002 | 81 |
|
2052 | 82 |
bind_thm ("OR2_parts_sees_Spy", |
83 |
OR2_analz_sees_Spy RS (impOfSubs analz_subset_parts)); |
|
84 |
bind_thm ("OR4_parts_sees_Spy", |
|
85 |
OR4_analz_sees_Spy RS (impOfSubs analz_subset_parts)); |
|
86 |
||
2002 | 87 |
val parts_Fake_tac = |
2052 | 88 |
forward_tac [OR2_parts_sees_Spy] 4 THEN |
89 |
forward_tac [OR4_parts_sees_Spy] 6 THEN |
|
2131 | 90 |
forward_tac [Oops_parts_sees_Spy] 7; |
91 |
||
92 |
(*For proving the easier theorems about X ~: parts (sees lost Spy evs) *) |
|
93 |
fun parts_induct_tac i = SELECT_GOAL |
|
94 |
(DETERM (etac otway.induct 1 THEN parts_Fake_tac THEN |
|
95 |
(*Fake message*) |
|
96 |
TRY (best_tac (!claset addDs [impOfSubs analz_subset_parts, |
|
97 |
impOfSubs Fake_parts_insert] |
|
98 |
addss (!simpset)) 2)) THEN |
|
99 |
(*Base case*) |
|
100 |
fast_tac (!claset addss (!simpset)) 1 THEN |
|
101 |
ALLGOALS Asm_simp_tac) i; |
|
2002 | 102 |
|
103 |
||
2052 | 104 |
(** Theorems of the form X ~: parts (sees lost Spy evs) imply that NOBODY |
2002 | 105 |
sends messages containing X! **) |
106 |
||
2131 | 107 |
(*Spy never sees another agent's shared key! (unless it's lost at start)*) |
2002 | 108 |
goal thy |
2131 | 109 |
"!!evs. evs : otway \ |
110 |
\ ==> (Key (shrK A) : parts (sees lost Spy evs)) = (A : lost)"; |
|
111 |
by (parts_induct_tac 1); |
|
2002 | 112 |
by (Auto_tac()); |
2131 | 113 |
qed "Spy_see_shrK"; |
114 |
Addsimps [Spy_see_shrK]; |
|
2002 | 115 |
|
2131 | 116 |
goal thy |
117 |
"!!evs. evs : otway \ |
|
118 |
\ ==> (Key (shrK A) : analz (sees lost Spy evs)) = (A : lost)"; |
|
119 |
by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset)); |
|
120 |
qed "Spy_analz_shrK"; |
|
121 |
Addsimps [Spy_analz_shrK]; |
|
2002 | 122 |
|
2131 | 123 |
goal thy "!!A. [| Key (shrK A) : parts (sees lost Spy evs); \ |
124 |
\ evs : otway |] ==> A:lost"; |
|
125 |
by (fast_tac (!claset addDs [Spy_see_shrK]) 1); |
|
126 |
qed "Spy_see_shrK_D"; |
|
2002 | 127 |
|
2131 | 128 |
bind_thm ("Spy_analz_shrK_D", analz_subset_parts RS subsetD RS Spy_see_shrK_D); |
129 |
AddSDs [Spy_see_shrK_D, Spy_analz_shrK_D]; |
|
2002 | 130 |
|
131 |
||
132 |
(*** Future keys can't be seen or used! ***) |
|
133 |
||
2160 | 134 |
(*Nobody can have SEEN keys that will be generated in the future. *) |
135 |
goal thy "!!evs. evs : otway ==> \ |
|
2002 | 136 |
\ length evs <= length evs' --> \ |
2160 | 137 |
\ Key (newK evs') ~: parts (sees lost Spy evs)"; |
2131 | 138 |
by (parts_induct_tac 1); |
2264
f298678bd54a
Weaking of injectivity assumptions for newK and newN:
paulson
parents:
2166
diff
changeset
|
139 |
by (REPEAT_FIRST (best_tac (!claset addEs [leD RS notE] |
f298678bd54a
Weaking of injectivity assumptions for newK and newN:
paulson
parents:
2166
diff
changeset
|
140 |
addDs [impOfSubs analz_subset_parts, |
2160 | 141 |
impOfSubs parts_insert_subset_Un, |
142 |
Suc_leD] |
|
2032 | 143 |
addss (!simpset)))); |
2160 | 144 |
qed_spec_mp "new_keys_not_seen"; |
2002 | 145 |
Addsimps [new_keys_not_seen]; |
146 |
||
2160 | 147 |
(*Variant: old messages must contain old keys!*) |
2002 | 148 |
goal thy |
149 |
"!!evs. [| Says A B X : set_of_list evs; \ |
|
150 |
\ Key (newK evt) : parts {X}; \ |
|
151 |
\ evs : otway \ |
|
152 |
\ |] ==> length evt < length evs"; |
|
2032 | 153 |
by (rtac ccontr 1); |
2052 | 154 |
by (dtac leI 1); |
2032 | 155 |
by (fast_tac (!claset addSDs [new_keys_not_seen, Says_imp_sees_Spy] |
2052 | 156 |
addIs [impOfSubs parts_mono]) 1); |
2002 | 157 |
qed "Says_imp_old_keys"; |
158 |
||
159 |
||
2160 | 160 |
(*** Future nonces can't be seen or used! ***) |
2002 | 161 |
|
162 |
goal thy "!!evs. evs : otway ==> \ |
|
163 |
\ length evs <= length evs' --> \ |
|
2160 | 164 |
\ Nonce (newN evs') ~: parts (sees lost Spy evs)"; |
165 |
by (parts_induct_tac 1); |
|
166 |
by (REPEAT_FIRST |
|
167 |
(fast_tac (!claset addSEs partsEs |
|
168 |
addSDs [Says_imp_sees_Spy RS parts.Inj] |
|
2264
f298678bd54a
Weaking of injectivity assumptions for newK and newN:
paulson
parents:
2166
diff
changeset
|
169 |
addEs [leD RS notE] |
2160 | 170 |
addDs [impOfSubs analz_subset_parts, |
171 |
impOfSubs parts_insert_subset_Un, Suc_leD] |
|
172 |
addss (!simpset)))); |
|
173 |
qed_spec_mp "new_nonces_not_seen"; |
|
2002 | 174 |
Addsimps [new_nonces_not_seen]; |
175 |
||
2160 | 176 |
(*Variant: old messages must contain old nonces!*) |
2002 | 177 |
goal thy |
2160 | 178 |
"!!evs. [| Says A B X : set_of_list evs; \ |
2002 | 179 |
\ Nonce (newN evt) : parts {X}; \ |
180 |
\ evs : otway \ |
|
181 |
\ |] ==> length evt < length evs"; |
|
2032 | 182 |
by (rtac ccontr 1); |
2052 | 183 |
by (dtac leI 1); |
2032 | 184 |
by (fast_tac (!claset addSDs [new_nonces_not_seen, Says_imp_sees_Spy] |
2052 | 185 |
addIs [impOfSubs parts_mono]) 1); |
2002 | 186 |
qed "Says_imp_old_nonces"; |
187 |
||
188 |
||
189 |
(*Nobody can have USED keys that will be generated in the future. |
|
190 |
...very like new_keys_not_seen*) |
|
191 |
goal thy "!!evs. evs : otway ==> \ |
|
192 |
\ length evs <= length evs' --> \ |
|
2160 | 193 |
\ newK evs' ~: keysFor (parts (sees lost Spy evs))"; |
2131 | 194 |
by (parts_induct_tac 1); |
2002 | 195 |
(*OR1 and OR3*) |
196 |
by (EVERY (map (fast_tac (!claset addDs [Suc_leD] addss (!simpset))) [4,2])); |
|
197 |
(*Fake, OR2, OR4: these messages send unknown (X) components*) |
|
2107
23e8f15ec95f
The new proof of the lemma for new_nonces_not_seen is faster
paulson
parents:
2052
diff
changeset
|
198 |
by (REPEAT |
23e8f15ec95f
The new proof of the lemma for new_nonces_not_seen is faster
paulson
parents:
2052
diff
changeset
|
199 |
(best_tac |
2002 | 200 |
(!claset addDs [impOfSubs (analz_subset_parts RS keysFor_mono), |
2032 | 201 |
impOfSubs (parts_insert_subset_Un RS keysFor_mono), |
202 |
Suc_leD] |
|
203 |
addEs [new_keys_not_seen RS not_parts_not_analz RSN(2,rev_notE)] |
|
2107
23e8f15ec95f
The new proof of the lemma for new_nonces_not_seen is faster
paulson
parents:
2052
diff
changeset
|
204 |
addss (!simpset)) 1)); |
2160 | 205 |
qed_spec_mp "new_keys_not_used"; |
2002 | 206 |
|
207 |
bind_thm ("new_keys_not_analzd", |
|
2032 | 208 |
[analz_subset_parts RS keysFor_mono, |
209 |
new_keys_not_used] MRS contra_subsetD); |
|
2002 | 210 |
|
211 |
Addsimps [new_keys_not_used, new_keys_not_analzd]; |
|
212 |
||
213 |
||
2131 | 214 |
|
215 |
(*** Proofs involving analz ***) |
|
216 |
||
217 |
(*Describes the form of K and NA when the Server sends this message. Also |
|
218 |
for Oops case.*) |
|
219 |
goal thy |
|
220 |
"!!evs. [| Says Server B \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
221 |
\ {|NA, X, Crypt (shrK B) {|NB, K|}|} : set_of_list evs; \ |
2131 | 222 |
\ evs : otway |] \ |
223 |
\ ==> (EX evt: otway. K = Key(newK evt)) & \ |
|
224 |
\ (EX i. NA = Nonce i) & (EX j. NB = Nonce j)"; |
|
225 |
by (etac rev_mp 1); |
|
226 |
by (etac otway.induct 1); |
|
227 |
by (ALLGOALS (fast_tac (!claset addss (!simpset)))); |
|
228 |
qed "Says_Server_message_form"; |
|
229 |
||
230 |
||
231 |
(*For proofs involving analz.*) |
|
232 |
val analz_Fake_tac = |
|
233 |
dtac OR2_analz_sees_Spy 4 THEN |
|
234 |
dtac OR4_analz_sees_Spy 6 THEN |
|
235 |
forward_tac [Says_Server_message_form] 7 THEN |
|
236 |
assume_tac 7 THEN Full_simp_tac 7 THEN |
|
237 |
REPEAT ((eresolve_tac [bexE, exE, conjE] ORELSE' hyp_subst_tac) 7); |
|
2002 | 238 |
|
239 |
||
240 |
(**** |
|
241 |
The following is to prove theorems of the form |
|
242 |
||
2052 | 243 |
Key K : analz (insert (Key (newK evt)) (sees lost Spy evs)) ==> |
244 |
Key K : analz (sees lost Spy evs) |
|
2002 | 245 |
|
246 |
A more general formula must be proved inductively. |
|
247 |
||
248 |
****) |
|
249 |
||
250 |
||
251 |
(** Session keys are not used to encrypt other session keys **) |
|
252 |
||
253 |
(*Lemma for the trivial direction of the if-and-only-if*) |
|
254 |
goal thy |
|
255 |
"!!evs. (Key K : analz (Key``nE Un sEe)) --> \ |
|
256 |
\ (K : nE | Key K : analz sEe) ==> \ |
|
257 |
\ (Key K : analz (Key``nE Un sEe)) = (K : nE | Key K : analz sEe)"; |
|
258 |
by (fast_tac (!claset addSEs [impOfSubs analz_mono]) 1); |
|
259 |
val lemma = result(); |
|
260 |
||
261 |
||
262 |
(*The equality makes the induction hypothesis easier to apply*) |
|
263 |
goal thy |
|
264 |
"!!evs. evs : otway ==> \ |
|
2052 | 265 |
\ ALL K E. (Key K : analz (Key``(newK``E) Un (sees lost Spy evs))) = \ |
266 |
\ (K : newK``E | Key K : analz (sees lost Spy evs))"; |
|
2032 | 267 |
by (etac otway.induct 1); |
2131 | 268 |
by analz_Fake_tac; |
2002 | 269 |
by (REPEAT_FIRST (ares_tac [allI, lemma])); |
2131 | 270 |
by (ALLGOALS |
2002 | 271 |
(asm_simp_tac |
272 |
(!simpset addsimps ([insert_Key_singleton, insert_Key_image, pushKey_newK] |
|
2032 | 273 |
@ pushes) |
2002 | 274 |
setloop split_tac [expand_if]))); |
275 |
(** LEVEL 7 **) |
|
2131 | 276 |
(*OR4, OR2, Fake*) |
277 |
by (EVERY (map spy_analz_tac [5,3,2])); |
|
278 |
(*Oops, OR3, Base*) |
|
279 |
by (REPEAT (fast_tac (!claset addIs [image_eqI] addss (!simpset)) 1)); |
|
2002 | 280 |
qed_spec_mp "analz_image_newK"; |
281 |
||
282 |
||
283 |
goal thy |
|
284 |
"!!evs. evs : otway ==> \ |
|
2052 | 285 |
\ Key K : analz (insert (Key (newK evt)) (sees lost Spy evs)) = \ |
286 |
\ (K = newK evt | Key K : analz (sees lost Spy evs))"; |
|
2002 | 287 |
by (asm_simp_tac (HOL_ss addsimps [pushKey_newK, analz_image_newK, |
2032 | 288 |
insert_Key_singleton]) 1); |
2002 | 289 |
by (Fast_tac 1); |
290 |
qed "analz_insert_Key_newK"; |
|
291 |
||
292 |
||
2131 | 293 |
(*** The Key K uniquely identifies the Server's message. **) |
2002 | 294 |
|
295 |
goal thy |
|
2131 | 296 |
"!!evs. evs : otway ==> \ |
297 |
\ EX B' NA' NB' X'. ALL B NA NB X. \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
298 |
\ Says Server B {|NA, X, Crypt (shrK B) {|NB, K|}|} : set_of_list evs --> \ |
2131 | 299 |
\ B=B' & NA=NA' & NB=NB' & X=X'"; |
2032 | 300 |
by (etac otway.induct 1); |
2002 | 301 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [all_conj_distrib]))); |
302 |
by (Step_tac 1); |
|
303 |
(*Remaining cases: OR3 and OR4*) |
|
304 |
by (ex_strip_tac 2); |
|
305 |
by (Fast_tac 2); |
|
2107
23e8f15ec95f
The new proof of the lemma for new_nonces_not_seen is faster
paulson
parents:
2052
diff
changeset
|
306 |
by (expand_case_tac "K = ?y" 1); |
23e8f15ec95f
The new proof of the lemma for new_nonces_not_seen is faster
paulson
parents:
2052
diff
changeset
|
307 |
by (REPEAT (ares_tac [refl,exI,impI,conjI] 2)); |
2002 | 308 |
(*...we assume X is a very new message, and handle this case by contradiction*) |
309 |
by (fast_tac (!claset addEs [Says_imp_old_keys RS less_irrefl] |
|
2032 | 310 |
delrules [conjI] (*prevent split-up into 4 subgoals*) |
311 |
addss (!simpset addsimps [parts_insertI])) 1); |
|
2002 | 312 |
val lemma = result(); |
313 |
||
314 |
goal thy |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
315 |
"!!evs. [| Says Server B {|NA, X, Crypt (shrK B) {|NB, K|}|} \ |
2002 | 316 |
\ : set_of_list evs; \ |
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
317 |
\ Says Server B' {|NA',X',Crypt (shrK B') {|NB',K|}|} \ |
2002 | 318 |
\ : set_of_list evs; \ |
2131 | 319 |
\ evs : otway |] ==> X=X' & B=B' & NA=NA' & NB=NB'"; |
2032 | 320 |
by (dtac lemma 1); |
2002 | 321 |
by (REPEAT (etac exE 1)); |
322 |
(*Duplicate the assumption*) |
|
323 |
by (forw_inst_tac [("psi", "ALL C.?P(C)")] asm_rl 1); |
|
324 |
by (fast_tac (!claset addSDs [spec]) 1); |
|
325 |
qed "unique_session_keys"; |
|
326 |
||
327 |
||
2131 | 328 |
(*Crucial security property, but not itself enough to guarantee correctness!*) |
329 |
goal thy |
|
2166 | 330 |
"!!evs. [| A ~: lost; B ~: lost; evs : otway |] \ |
331 |
\ ==> Says Server B \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
332 |
\ {|NA, Crypt (shrK A) {|NA, Key K|}, \ |
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
333 |
\ Crypt (shrK B) {|NB, Key K|}|} : set_of_list evs --> \ |
2166 | 334 |
\ Says B Spy {|NA, NB, Key K|} ~: set_of_list evs --> \ |
335 |
\ Key K ~: analz (sees lost Spy evs)"; |
|
2131 | 336 |
by (etac otway.induct 1); |
337 |
by analz_Fake_tac; |
|
338 |
by (ALLGOALS |
|
339 |
(asm_full_simp_tac |
|
340 |
(!simpset addsimps ([analz_subset_parts RS contra_subsetD, |
|
341 |
analz_insert_Key_newK] @ pushes) |
|
342 |
setloop split_tac [expand_if]))); |
|
343 |
(*OR3*) |
|
2166 | 344 |
by (fast_tac (!claset addEs [Says_imp_old_keys RS less_irrefl] |
2131 | 345 |
addss (!simpset addsimps [parts_insert2])) 3); |
346 |
(*OR4, OR2, Fake*) |
|
347 |
by (REPEAT_FIRST (resolve_tac [conjI, impI] ORELSE' spy_analz_tac)); |
|
348 |
(*Oops*) (** LEVEL 5 **) |
|
349 |
by (fast_tac (!claset delrules [disjE] |
|
350 |
addDs [unique_session_keys] addss (!simpset)) 1); |
|
351 |
val lemma = result() RS mp RS mp RSN(2,rev_notE); |
|
352 |
||
353 |
||
354 |
goal thy |
|
355 |
"!!evs. [| Says Server B \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
356 |
\ {|NA, Crypt (shrK A) {|NA, K|}, \ |
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
357 |
\ Crypt (shrK B) {|NB, K|}|} : set_of_list evs; \ |
2131 | 358 |
\ Says B Spy {|NA, NB, K|} ~: set_of_list evs; \ |
359 |
\ A ~: lost; B ~: lost; evs : otway |] \ |
|
360 |
\ ==> K ~: analz (sees lost Spy evs)"; |
|
361 |
by (forward_tac [Says_Server_message_form] 1 THEN assume_tac 1); |
|
362 |
by (fast_tac (!claset addSEs [lemma]) 1); |
|
363 |
qed "Spy_not_see_encrypted_key"; |
|
364 |
||
365 |
||
366 |
(*** Attempting to prove stronger properties ***) |
|
367 |
||
2052 | 368 |
(*Only OR1 can have caused such a part of a message to appear. |
369 |
I'm not sure why A ~= B premise is needed: OtwayRees.ML doesn't need it. |
|
370 |
Perhaps it's because OR2 has two similar-looking encrypted messages in |
|
371 |
this version.*) |
|
2002 | 372 |
goal thy |
2131 | 373 |
"!!evs. [| A ~: lost; A ~= B; evs : otway |] \ |
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
374 |
\ ==> Crypt (shrK A) {|NA, Agent A, Agent B|} \ |
2052 | 375 |
\ : parts (sees lost Spy evs) --> \ |
2131 | 376 |
\ Says A B {|NA, Agent A, Agent B, \ |
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
377 |
\ Crypt (shrK A) {|NA, Agent A, Agent B|}|} \ |
2002 | 378 |
\ : set_of_list evs"; |
2131 | 379 |
by (parts_induct_tac 1); |
380 |
by (Fast_tac 1); |
|
2002 | 381 |
qed_spec_mp "Crypt_imp_OR1"; |
382 |
||
383 |
||
2131 | 384 |
(*Crucial property: If the encrypted message appears, and A has used NA |
385 |
to start a run, then it originated with the Server!*) |
|
386 |
(*Only it is FALSE. Somebody could make a fake message to Server |
|
2002 | 387 |
substituting some other nonce NA' for NB.*) |
388 |
goal thy |
|
2052 | 389 |
"!!evs. [| A ~: lost; A ~= Spy; evs : otway |] \ |
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
390 |
\ ==> Crypt (shrK A) {|NA, Key K|} : parts (sees lost Spy evs) --> \ |
2131 | 391 |
\ Says A B {|NA, Agent A, Agent B, \ |
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
392 |
\ Crypt (shrK A) {|NA, Agent A, Agent B|}|} \ |
2131 | 393 |
\ : set_of_list evs --> \ |
394 |
\ (EX B NB. Says Server B \ |
|
395 |
\ {|NA, \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
396 |
\ Crypt (shrK A) {|NA, Key K|}, \ |
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
397 |
\ Crypt (shrK B) {|NB, Key K|}|} \ |
2002 | 398 |
\ : set_of_list evs)"; |
2131 | 399 |
by (parts_induct_tac 1); |
2002 | 400 |
(*OR1: it cannot be a new Nonce, contradiction.*) |
401 |
by (fast_tac (!claset addSIs [parts_insertI] |
|
2032 | 402 |
addSEs partsEs |
403 |
addEs [Says_imp_old_nonces RS less_irrefl] |
|
404 |
addss (!simpset)) 1); |
|
2002 | 405 |
(*OR4*) |
406 |
by (REPEAT (Safe_step_tac 2)); |
|
2052 | 407 |
by (REPEAT (best_tac (!claset addSDs [parts_cut]) 3)); |
408 |
by (fast_tac (!claset addSIs [Crypt_imp_OR1] |
|
409 |
addEs partsEs |
|
410 |
addDs [Says_imp_sees_Spy RS parts.Inj]) 2); |
|
2131 | 411 |
(*OR3*) (** LEVEL 5 **) |
2002 | 412 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [ex_disj_distrib]))); |
2052 | 413 |
by (step_tac (!claset delrules [disjCI, impCE]) 1); |
2002 | 414 |
(*The hypotheses at this point suggest an attack in which nonce NA is used |
2052 | 415 |
in two different roles: |
416 |
Says B' Server |
|
417 |
{|Nonce NAa, Agent Aa, Agent A, |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
418 |
Crypt (shrK Aa) {|Nonce NAa, Agent Aa, Agent A|}, Nonce NA, |
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
419 |
Crypt (shrK A) {|Nonce NAa, Agent Aa, Agent A|}|} |
2052 | 420 |
: set_of_list evsa; |
421 |
Says A B |
|
422 |
{|Nonce NA, Agent A, Agent B, |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2264
diff
changeset
|
423 |
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|}|} |
2052 | 424 |
: set_of_list evsa |
425 |
*) |
|
2131 | 426 |
writeln "GIVE UP! on NA_Crypt_imp_Server_msg"; |
2002 | 427 |
|
428 |
||
2052 | 429 |
(*Thus the key property A_can_trust probably fails too.*) |