paulson@2002
|
1 |
(* Title: HOL/Auth/OtwayRees_Bad
|
paulson@2002
|
2 |
ID: $Id$
|
paulson@2002
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
paulson@2002
|
4 |
Copyright 1996 University of Cambridge
|
paulson@14207
|
5 |
*)
|
paulson@2002
|
6 |
|
paulson@14207
|
7 |
|
paulson@14207
|
8 |
header{*The Otway-Rees Protocol: The Faulty BAN Version*}
|
paulson@2002
|
9 |
|
haftmann@16417
|
10 |
theory OtwayRees_Bad imports Public begin
|
paulson@14207
|
11 |
|
paulson@14207
|
12 |
text{*The FAULTY version omitting encryption of Nonce NB, as suggested on
|
paulson@14207
|
13 |
page 247 of
|
paulson@14207
|
14 |
Burrows, Abadi and Needham (1988). A Logic of Authentication.
|
paulson@14207
|
15 |
Proc. Royal Soc. 426
|
paulson@11251
|
16 |
|
paulson@11251
|
17 |
This file illustrates the consequences of such errors. We can still prove
|
paulson@14207
|
18 |
impressive-looking properties such as @{text Spy_not_see_encrypted_key}, yet
|
paulson@14207
|
19 |
the protocol is open to a middleperson attack. Attempting to prove some key
|
paulson@14207
|
20 |
lemmas indicates the possibility of this attack.*}
|
paulson@2052
|
21 |
|
paulson@11251
|
22 |
consts otway :: "event list set"
|
paulson@11251
|
23 |
inductive "otway"
|
paulson@11251
|
24 |
intros
|
paulson@14225
|
25 |
Nil: --{*The empty trace*}
|
paulson@14225
|
26 |
"[] \<in> otway"
|
paulson@2002
|
27 |
|
paulson@14225
|
28 |
Fake: --{*The Spy may say anything he can say. The sender field is correct,
|
paulson@14225
|
29 |
but agents don't use that information.*}
|
paulson@14225
|
30 |
"[| evsf \<in> otway; X \<in> synth (analz (knows Spy evsf)) |]
|
paulson@11251
|
31 |
==> Says Spy B X # evsf \<in> otway"
|
paulson@2002
|
32 |
|
paulson@14225
|
33 |
|
paulson@14225
|
34 |
Reception: --{*A message that has been sent can be received by the
|
paulson@14225
|
35 |
intended recipient.*}
|
paulson@14225
|
36 |
"[| evsr \<in> otway; Says A B X \<in>set evsr |]
|
paulson@11251
|
37 |
==> Gets B X # evsr \<in> otway"
|
paulson@6308
|
38 |
|
paulson@14225
|
39 |
OR1: --{*Alice initiates a protocol run*}
|
paulson@14225
|
40 |
"[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
|
paulson@11251
|
41 |
==> Says A B {|Nonce NA, Agent A, Agent B,
|
paulson@11251
|
42 |
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
|
paulson@11251
|
43 |
# evs1 \<in> otway"
|
paulson@2002
|
44 |
|
paulson@14225
|
45 |
OR2: --{*Bob's response to Alice's message.
|
paulson@14225
|
46 |
This variant of the protocol does NOT encrypt NB.*}
|
paulson@14225
|
47 |
"[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
|
paulson@11251
|
48 |
Gets B {|Nonce NA, Agent A, Agent B, X|} \<in> set evs2 |]
|
paulson@11251
|
49 |
==> Says B Server
|
paulson@2516
|
50 |
{|Nonce NA, Agent A, Agent B, X, Nonce NB,
|
paulson@2284
|
51 |
Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
|
paulson@11251
|
52 |
# evs2 \<in> otway"
|
paulson@2002
|
53 |
|
paulson@14225
|
54 |
OR3: --{*The Server receives Bob's message and checks that the three NAs
|
paulson@2002
|
55 |
match. Then he sends a new session key to Bob with a packet for
|
paulson@14225
|
56 |
forwarding to Alice.*}
|
paulson@14225
|
57 |
"[| evs3 \<in> otway; Key KAB \<notin> used evs3;
|
paulson@11251
|
58 |
Gets Server
|
paulson@11251
|
59 |
{|Nonce NA, Agent A, Agent B,
|
paulson@11251
|
60 |
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|},
|
paulson@11251
|
61 |
Nonce NB,
|
paulson@2284
|
62 |
Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
|
paulson@11251
|
63 |
\<in> set evs3 |]
|
paulson@11251
|
64 |
==> Says Server B
|
paulson@11251
|
65 |
{|Nonce NA,
|
paulson@2516
|
66 |
Crypt (shrK A) {|Nonce NA, Key KAB|},
|
paulson@2516
|
67 |
Crypt (shrK B) {|Nonce NB, Key KAB|}|}
|
paulson@11251
|
68 |
# evs3 \<in> otway"
|
paulson@2002
|
69 |
|
paulson@14225
|
70 |
OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
|
paulson@14225
|
71 |
those in the message he previously sent the Server.
|
paulson@14225
|
72 |
Need @{term "B \<noteq> Server"} because we allow messages to self.*}
|
paulson@14225
|
73 |
"[| evs4 \<in> otway; B \<noteq> Server;
|
paulson@4522
|
74 |
Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB,
|
paulson@4522
|
75 |
Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
|
paulson@11251
|
76 |
\<in> set evs4;
|
paulson@6308
|
77 |
Gets B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
|
paulson@11251
|
78 |
\<in> set evs4 |]
|
paulson@11251
|
79 |
==> Says B A {|Nonce NA, X|} # evs4 \<in> otway"
|
paulson@2002
|
80 |
|
paulson@14225
|
81 |
Oops: --{*This message models possible leaks of session keys. The nonces
|
paulson@14225
|
82 |
identify the protocol run.*}
|
paulson@14225
|
83 |
"[| evso \<in> otway;
|
paulson@2284
|
84 |
Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
|
paulson@11251
|
85 |
\<in> set evso |]
|
paulson@11251
|
86 |
==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> otway"
|
paulson@11251
|
87 |
|
paulson@11251
|
88 |
|
paulson@11251
|
89 |
declare Says_imp_knows_Spy [THEN analz.Inj, dest]
|
paulson@11251
|
90 |
declare parts.Body [dest]
|
paulson@11251
|
91 |
declare analz_into_parts [dest]
|
paulson@11251
|
92 |
declare Fake_parts_insert_in_Un [dest]
|
paulson@11251
|
93 |
|
paulson@14225
|
94 |
text{*A "possibility property": there are traces that reach the end*}
|
paulson@14200
|
95 |
lemma "[| B \<noteq> Server; Key K \<notin> used [] |]
|
paulson@14200
|
96 |
==> \<exists>NA. \<exists>evs \<in> otway.
|
paulson@11251
|
97 |
Says B A {|Nonce NA, Crypt (shrK A) {|Nonce NA, Key K|}|}
|
paulson@11251
|
98 |
\<in> set evs"
|
paulson@11251
|
99 |
apply (intro exI bexI)
|
paulson@11251
|
100 |
apply (rule_tac [2] otway.Nil
|
paulson@11251
|
101 |
[THEN otway.OR1, THEN otway.Reception,
|
paulson@11251
|
102 |
THEN otway.OR2, THEN otway.Reception,
|
paulson@14200
|
103 |
THEN otway.OR3, THEN otway.Reception, THEN otway.OR4])
|
paulson@14200
|
104 |
apply (possibility, simp add: used_Cons)
|
paulson@11251
|
105 |
done
|
paulson@11251
|
106 |
|
paulson@11251
|
107 |
lemma Gets_imp_Says [dest!]:
|
paulson@11251
|
108 |
"[| Gets B X \<in> set evs; evs \<in> otway |] ==> \<exists>A. Says A B X \<in> set evs"
|
paulson@11251
|
109 |
apply (erule rev_mp)
|
paulson@13507
|
110 |
apply (erule otway.induct, auto)
|
paulson@11251
|
111 |
done
|
paulson@11251
|
112 |
|
paulson@11251
|
113 |
|
paulson@14225
|
114 |
subsection{*For reasoning about the encrypted portion of messages *}
|
paulson@11251
|
115 |
|
paulson@11251
|
116 |
lemma OR2_analz_knows_Spy:
|
paulson@11251
|
117 |
"[| Gets B {|N, Agent A, Agent B, X|} \<in> set evs; evs \<in> otway |]
|
paulson@11251
|
118 |
==> X \<in> analz (knows Spy evs)"
|
paulson@11251
|
119 |
by blast
|
paulson@11251
|
120 |
|
paulson@11251
|
121 |
lemma OR4_analz_knows_Spy:
|
paulson@11251
|
122 |
"[| Gets B {|N, X, Crypt (shrK B) X'|} \<in> set evs; evs \<in> otway |]
|
paulson@11251
|
123 |
==> X \<in> analz (knows Spy evs)"
|
paulson@11251
|
124 |
by blast
|
paulson@11251
|
125 |
|
paulson@11251
|
126 |
lemma Oops_parts_knows_Spy:
|
paulson@11251
|
127 |
"Says Server B {|NA, X, Crypt K' {|NB,K|}|} \<in> set evs
|
paulson@11251
|
128 |
==> K \<in> parts (knows Spy evs)"
|
paulson@11251
|
129 |
by blast
|
paulson@11251
|
130 |
|
paulson@14225
|
131 |
text{*Forwarding lemma: see comments in OtwayRees.thy*}
|
paulson@11251
|
132 |
lemmas OR2_parts_knows_Spy =
|
paulson@11251
|
133 |
OR2_analz_knows_Spy [THEN analz_into_parts, standard]
|
paulson@11251
|
134 |
|
paulson@11251
|
135 |
|
paulson@14225
|
136 |
text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
|
paulson@14225
|
137 |
NOBODY sends messages containing X! *}
|
paulson@11251
|
138 |
|
paulson@14225
|
139 |
text{*Spy never sees a good agent's shared key!*}
|
paulson@11251
|
140 |
lemma Spy_see_shrK [simp]:
|
paulson@11251
|
141 |
"evs \<in> otway ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
|
paulson@14225
|
142 |
by (erule otway.induct, force,
|
paulson@14225
|
143 |
drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+)
|
paulson@14225
|
144 |
|
paulson@11251
|
145 |
|
paulson@11251
|
146 |
lemma Spy_analz_shrK [simp]:
|
paulson@11251
|
147 |
"evs \<in> otway ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
|
paulson@11251
|
148 |
by auto
|
paulson@11251
|
149 |
|
paulson@11251
|
150 |
lemma Spy_see_shrK_D [dest!]:
|
paulson@11251
|
151 |
"[|Key (shrK A) \<in> parts (knows Spy evs); evs \<in> otway|] ==> A \<in> bad"
|
paulson@11251
|
152 |
by (blast dest: Spy_see_shrK)
|
paulson@11251
|
153 |
|
paulson@11251
|
154 |
|
paulson@14225
|
155 |
subsection{*Proofs involving analz *}
|
paulson@11251
|
156 |
|
paulson@14225
|
157 |
text{*Describes the form of K and NA when the Server sends this message. Also
|
paulson@14225
|
158 |
for Oops case.*}
|
paulson@11251
|
159 |
lemma Says_Server_message_form:
|
paulson@11251
|
160 |
"[| Says Server B {|NA, X, Crypt (shrK B) {|NB, Key K|}|} \<in> set evs;
|
paulson@11251
|
161 |
evs \<in> otway |]
|
paulson@11251
|
162 |
==> K \<notin> range shrK & (\<exists>i. NA = Nonce i) & (\<exists>j. NB = Nonce j)"
|
paulson@11251
|
163 |
apply (erule rev_mp)
|
nipkow@17778
|
164 |
apply (erule otway.induct, simp_all)
|
paulson@11251
|
165 |
done
|
paulson@11251
|
166 |
|
paulson@11251
|
167 |
|
paulson@11251
|
168 |
(****
|
paulson@11251
|
169 |
The following is to prove theorems of the form
|
paulson@11251
|
170 |
|
paulson@11251
|
171 |
Key K \<in> analz (insert (Key KAB) (knows Spy evs)) ==>
|
paulson@11251
|
172 |
Key K \<in> analz (knows Spy evs)
|
paulson@11251
|
173 |
|
paulson@11251
|
174 |
A more general formula must be proved inductively.
|
paulson@11251
|
175 |
****)
|
paulson@11251
|
176 |
|
paulson@11251
|
177 |
|
paulson@14225
|
178 |
text{*Session keys are not used to encrypt other session keys*}
|
paulson@11251
|
179 |
|
paulson@14225
|
180 |
text{*The equality makes the induction hypothesis easier to apply*}
|
paulson@11251
|
181 |
lemma analz_image_freshK [rule_format]:
|
paulson@11251
|
182 |
"evs \<in> otway ==>
|
paulson@11251
|
183 |
\<forall>K KK. KK <= -(range shrK) -->
|
paulson@11251
|
184 |
(Key K \<in> analz (Key`KK Un (knows Spy evs))) =
|
paulson@11251
|
185 |
(K \<in> KK | Key K \<in> analz (knows Spy evs))"
|
paulson@14207
|
186 |
apply (erule otway.induct)
|
paulson@14207
|
187 |
apply (frule_tac [8] Says_Server_message_form)
|
paulson@14207
|
188 |
apply (drule_tac [7] OR4_analz_knows_Spy)
|
paulson@14207
|
189 |
apply (drule_tac [5] OR2_analz_knows_Spy, analz_freshK, spy_analz, auto)
|
paulson@11251
|
190 |
done
|
paulson@11251
|
191 |
|
paulson@11251
|
192 |
lemma analz_insert_freshK:
|
paulson@11251
|
193 |
"[| evs \<in> otway; KAB \<notin> range shrK |] ==>
|
wenzelm@11655
|
194 |
(Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =
|
paulson@11251
|
195 |
(K = KAB | Key K \<in> analz (knows Spy evs))"
|
paulson@11251
|
196 |
by (simp only: analz_image_freshK analz_image_freshK_simps)
|
paulson@11251
|
197 |
|
paulson@11251
|
198 |
|
paulson@14225
|
199 |
text{*The Key K uniquely identifies the Server's message. *}
|
paulson@11251
|
200 |
lemma unique_session_keys:
|
paulson@11251
|
201 |
"[| Says Server B {|NA, X, Crypt (shrK B) {|NB, K|}|} \<in> set evs;
|
paulson@11251
|
202 |
Says Server B' {|NA',X',Crypt (shrK B') {|NB',K|}|} \<in> set evs;
|
paulson@11251
|
203 |
evs \<in> otway |] ==> X=X' & B=B' & NA=NA' & NB=NB'"
|
paulson@11251
|
204 |
apply (erule rev_mp)
|
paulson@11251
|
205 |
apply (erule rev_mp)
|
paulson@11251
|
206 |
apply (erule otway.induct, simp_all)
|
paulson@14225
|
207 |
apply blast+ --{*OR3 and OR4*}
|
paulson@11251
|
208 |
done
|
paulson@11251
|
209 |
|
paulson@11251
|
210 |
|
paulson@14238
|
211 |
text{*Crucial secrecy property: Spy does not see the keys sent in msg OR3
|
paulson@11251
|
212 |
Does not in itself guarantee security: an attack could violate
|
paulson@14225
|
213 |
the premises, e.g. by having @{term "A=Spy"} *}
|
paulson@11251
|
214 |
lemma secrecy_lemma:
|
paulson@11251
|
215 |
"[| A \<notin> bad; B \<notin> bad; evs \<in> otway |]
|
paulson@11251
|
216 |
==> Says Server B
|
paulson@11251
|
217 |
{|NA, Crypt (shrK A) {|NA, Key K|},
|
paulson@11251
|
218 |
Crypt (shrK B) {|NB, Key K|}|} \<in> set evs -->
|
paulson@11251
|
219 |
Notes Spy {|NA, NB, Key K|} \<notin> set evs -->
|
paulson@11251
|
220 |
Key K \<notin> analz (knows Spy evs)"
|
paulson@11251
|
221 |
apply (erule otway.induct, force)
|
paulson@11251
|
222 |
apply (frule_tac [7] Says_Server_message_form)
|
paulson@11251
|
223 |
apply (drule_tac [6] OR4_analz_knows_Spy)
|
paulson@11251
|
224 |
apply (drule_tac [4] OR2_analz_knows_Spy)
|
paulson@14225
|
225 |
apply (simp_all add: analz_insert_eq analz_insert_freshK pushes)
|
paulson@14225
|
226 |
apply spy_analz --{*Fake*}
|
paulson@14225
|
227 |
apply (blast dest: unique_session_keys)+ --{*OR3, OR4, Oops*}
|
paulson@11251
|
228 |
done
|
paulson@11251
|
229 |
|
paulson@11251
|
230 |
|
paulson@11251
|
231 |
lemma Spy_not_see_encrypted_key:
|
paulson@11251
|
232 |
"[| Says Server B
|
paulson@11251
|
233 |
{|NA, Crypt (shrK A) {|NA, Key K|},
|
paulson@11251
|
234 |
Crypt (shrK B) {|NB, Key K|}|} \<in> set evs;
|
paulson@11251
|
235 |
Notes Spy {|NA, NB, Key K|} \<notin> set evs;
|
paulson@11251
|
236 |
A \<notin> bad; B \<notin> bad; evs \<in> otway |]
|
paulson@11251
|
237 |
==> Key K \<notin> analz (knows Spy evs)"
|
paulson@11251
|
238 |
by (blast dest: Says_Server_message_form secrecy_lemma)
|
paulson@11251
|
239 |
|
paulson@11251
|
240 |
|
paulson@14225
|
241 |
subsection{*Attempting to prove stronger properties *}
|
paulson@11251
|
242 |
|
paulson@14225
|
243 |
text{*Only OR1 can have caused such a part of a message to appear. The premise
|
paulson@14225
|
244 |
@{term "A \<noteq> B"} prevents OR2's similar-looking cryptogram from being picked
|
paulson@14225
|
245 |
up. Original Otway-Rees doesn't need it.*}
|
paulson@11251
|
246 |
lemma Crypt_imp_OR1 [rule_format]:
|
paulson@11251
|
247 |
"[| A \<notin> bad; A \<noteq> B; evs \<in> otway |]
|
paulson@11251
|
248 |
==> Crypt (shrK A) {|NA, Agent A, Agent B|} \<in> parts (knows Spy evs) -->
|
paulson@11251
|
249 |
Says A B {|NA, Agent A, Agent B,
|
paulson@11251
|
250 |
Crypt (shrK A) {|NA, Agent A, Agent B|}|} \<in> set evs"
|
paulson@14225
|
251 |
by (erule otway.induct, force,
|
paulson@14225
|
252 |
drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+)
|
paulson@11251
|
253 |
|
paulson@11251
|
254 |
|
paulson@14225
|
255 |
text{*Crucial property: If the encrypted message appears, and A has used NA
|
paulson@11251
|
256 |
to start a run, then it originated with the Server!
|
paulson@14225
|
257 |
The premise @{term "A \<noteq> B"} allows use of @{text Crypt_imp_OR1}*}
|
paulson@14225
|
258 |
text{*Only it is FALSE. Somebody could make a fake message to Server
|
paulson@14225
|
259 |
substituting some other nonce NA' for NB.*}
|
paulson@11251
|
260 |
lemma "[| A \<notin> bad; A \<noteq> B; evs \<in> otway |]
|
paulson@11251
|
261 |
==> Crypt (shrK A) {|NA, Key K|} \<in> parts (knows Spy evs) -->
|
paulson@11251
|
262 |
Says A B {|NA, Agent A, Agent B,
|
paulson@11251
|
263 |
Crypt (shrK A) {|NA, Agent A, Agent B|}|}
|
paulson@11251
|
264 |
\<in> set evs -->
|
paulson@11251
|
265 |
(\<exists>B NB. Says Server B
|
paulson@11251
|
266 |
{|NA,
|
paulson@11251
|
267 |
Crypt (shrK A) {|NA, Key K|},
|
paulson@11251
|
268 |
Crypt (shrK B) {|NB, Key K|}|} \<in> set evs)"
|
paulson@11251
|
269 |
apply (erule otway.induct, force,
|
paulson@13507
|
270 |
drule_tac [4] OR2_parts_knows_Spy, simp_all)
|
paulson@14225
|
271 |
apply blast --{*Fake*}
|
paulson@14225
|
272 |
apply blast --{*OR1: it cannot be a new Nonce, contradiction.*}
|
paulson@14225
|
273 |
txt{*OR3 and OR4*}
|
paulson@11251
|
274 |
apply (simp_all add: ex_disj_distrib)
|
paulson@14225
|
275 |
prefer 2 apply (blast intro!: Crypt_imp_OR1) --{*OR4*}
|
paulson@14225
|
276 |
txt{*OR3*}
|
paulson@11251
|
277 |
apply clarify
|
paulson@11251
|
278 |
(*The hypotheses at this point suggest an attack in which nonce NB is used
|
paulson@11251
|
279 |
in two different roles:
|
paulson@11251
|
280 |
Gets Server
|
paulson@11251
|
281 |
{|Nonce NA, Agent Aa, Agent A,
|
paulson@11251
|
282 |
Crypt (shrK Aa) {|Nonce NA, Agent Aa, Agent A|}, Nonce NB,
|
paulson@11251
|
283 |
Crypt (shrK A) {|Nonce NA, Agent Aa, Agent A|}|}
|
paulson@11251
|
284 |
\<in> set evs3
|
paulson@11251
|
285 |
Says A B
|
paulson@11251
|
286 |
{|Nonce NB, Agent A, Agent B,
|
paulson@11251
|
287 |
Crypt (shrK A) {|Nonce NB, Agent A, Agent B|}|}
|
paulson@11251
|
288 |
\<in> set evs3;
|
paulson@11251
|
289 |
*)
|
paulson@11251
|
290 |
|
paulson@11251
|
291 |
|
paulson@11251
|
292 |
(*Thus the key property A_can_trust probably fails too.*)
|
paulson@11251
|
293 |
oops
|
paulson@2002
|
294 |
|
paulson@2002
|
295 |
end
|