author | huffman |
Sun, 07 Mar 2010 16:39:31 -0800 | |
changeset 35642 | f478d5a9d238 |
parent 32960 | 69916a850301 |
child 37936 | 1e4c5015a72e |
permissions | -rw-r--r-- |
2090 | 1 |
(* Title: HOL/Auth/OtwayRees |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
5 |
*) |
2090 | 6 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
7 |
header{*The Otway-Rees Protocol as Modified by Abadi and Needham*} |
2090 | 8 |
|
16417 | 9 |
theory OtwayRees_AN imports Public begin |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
10 |
|
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
11 |
text{* |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
12 |
This simplified version has minimal encryption and explicit messages. |
2090 | 13 |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
14 |
Note that the formalization does not even assume that nonces are fresh. |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
15 |
This is because the protocol does not rely on uniqueness of nonces for |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
16 |
security, only for freshness, and the proof script does not prove freshness |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
17 |
properties. |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
18 |
|
2090 | 19 |
From page 11 of |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
20 |
Abadi and Needham (1996). |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
21 |
Prudent Engineering Practice for Cryptographic Protocols. |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
22 |
IEEE Trans. SE 22 (1) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
23 |
*} |
2090 | 24 |
|
23746 | 25 |
inductive_set otway :: "event list set" |
26 |
where |
|
14238 | 27 |
Nil: --{*The empty trace*} |
28 |
"[] \<in> otway" |
|
2090 | 29 |
|
23746 | 30 |
| Fake: --{*The Spy may say anything he can say. The sender field is correct, |
14238 | 31 |
but agents don't use that information.*} |
32 |
"[| evsf \<in> otway; X \<in> synth (analz (knows Spy evsf)) |] |
|
11251 | 33 |
==> Says Spy B X # evsf \<in> otway" |
2090 | 34 |
|
14238 | 35 |
|
23746 | 36 |
| Reception: --{*A message that has been sent can be received by the |
14238 | 37 |
intended recipient.*} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32366
diff
changeset
|
38 |
"[| evsr \<in> otway; Says A B X \<in>set evsr |] |
11251 | 39 |
==> Gets B X # evsr \<in> otway" |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5434
diff
changeset
|
40 |
|
23746 | 41 |
| OR1: --{*Alice initiates a protocol run*} |
14238 | 42 |
"evs1 \<in> otway |
11251 | 43 |
==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway" |
2090 | 44 |
|
23746 | 45 |
| OR2: --{*Bob's response to Alice's message.*} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32366
diff
changeset
|
46 |
"[| evs2 \<in> otway; |
11251 | 47 |
Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |] |
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
48 |
==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} |
11251 | 49 |
# evs2 \<in> otway" |
2090 | 50 |
|
23746 | 51 |
| OR3: --{*The Server receives Bob's message. Then he sends a new |
14238 | 52 |
session key to Bob with a packet for forwarding to Alice.*} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32366
diff
changeset
|
53 |
"[| evs3 \<in> otway; Key KAB \<notin> used evs3; |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5434
diff
changeset
|
54 |
Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|} |
11251 | 55 |
\<in>set evs3 |] |
56 |
==> Says Server B |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
57 |
{|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key KAB|}, |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
58 |
Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key KAB|}|} |
11251 | 59 |
# evs3 \<in> otway" |
2090 | 60 |
|
23746 | 61 |
| OR4: --{*Bob receives the Server's (?) message and compares the Nonces with |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32366
diff
changeset
|
62 |
those in the message he previously sent the Server. |
14238 | 63 |
Need @{term "B \<noteq> Server"} because we allow messages to self.*} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32366
diff
changeset
|
64 |
"[| evs4 \<in> otway; B \<noteq> Server; |
11251 | 65 |
Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \<in>set evs4; |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5434
diff
changeset
|
66 |
Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|} |
11251 | 67 |
\<in>set evs4 |] |
68 |
==> Says B A X # evs4 \<in> otway" |
|
2090 | 69 |
|
23746 | 70 |
| Oops: --{*This message models possible leaks of session keys. The nonces |
14238 | 71 |
identify the protocol run.*} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32366
diff
changeset
|
72 |
"[| evso \<in> otway; |
11251 | 73 |
Says Server B |
74 |
{|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|}, |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2131
diff
changeset
|
75 |
Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|} |
11251 | 76 |
\<in>set evso |] |
77 |
==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> otway" |
|
78 |
||
79 |
||
80 |
declare Says_imp_knows_Spy [THEN analz.Inj, dest] |
|
81 |
declare parts.Body [dest] |
|
82 |
declare analz_into_parts [dest] |
|
83 |
declare Fake_parts_insert_in_Un [dest] |
|
84 |
||
85 |
||
14238 | 86 |
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:
13507
diff
changeset
|
87 |
lemma "[| B \<noteq> Server; Key K \<notin> used [] |] |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13507
diff
changeset
|
88 |
==> \<exists>evs \<in> otway. |
11251 | 89 |
Says B A (Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|}) |
90 |
\<in> set evs" |
|
91 |
apply (intro exI bexI) |
|
92 |
apply (rule_tac [2] otway.Nil |
|
93 |
[THEN otway.OR1, THEN otway.Reception, |
|
94 |
THEN otway.OR2, THEN otway.Reception, |
|
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13507
diff
changeset
|
95 |
THEN otway.OR3, THEN otway.Reception, THEN otway.OR4]) |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13507
diff
changeset
|
96 |
apply (possibility, simp add: used_Cons) |
11251 | 97 |
done |
98 |
||
99 |
lemma Gets_imp_Says [dest!]: |
|
100 |
"[| Gets B X \<in> set evs; evs \<in> otway |] ==> \<exists>A. Says A B X \<in> set evs" |
|
101 |
by (erule rev_mp, erule otway.induct, auto) |
|
102 |
||
103 |
||
104 |
||
14238 | 105 |
text{* For reasoning about the encrypted portion of messages *} |
11251 | 106 |
|
107 |
lemma OR4_analz_knows_Spy: |
|
108 |
"[| Gets B {|X, Crypt(shrK B) X'|} \<in> set evs; evs \<in> otway |] |
|
109 |
==> X \<in> analz (knows Spy evs)" |
|
110 |
by blast |
|
111 |
||
112 |
||
14238 | 113 |
text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that |
114 |
NOBODY sends messages containing X! *} |
|
11251 | 115 |
|
14238 | 116 |
text{*Spy never sees a good agent's shared key!*} |
11251 | 117 |
lemma Spy_see_shrK [simp]: |
118 |
"evs \<in> otway ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)" |
|
14238 | 119 |
by (erule otway.induct, simp_all, blast+) |
11251 | 120 |
|
121 |
lemma Spy_analz_shrK [simp]: |
|
122 |
"evs \<in> otway ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)" |
|
123 |
by auto |
|
124 |
||
125 |
lemma Spy_see_shrK_D [dest!]: |
|
126 |
"[|Key (shrK A) \<in> parts (knows Spy evs); evs \<in> otway|] ==> A \<in> bad" |
|
127 |
by (blast dest: Spy_see_shrK) |
|
128 |
||
129 |
||
14238 | 130 |
subsection{*Proofs involving analz *} |
11251 | 131 |
|
14238 | 132 |
text{*Describes the form of K and NA when the Server sends this message.*} |
11251 | 133 |
lemma Says_Server_message_form: |
134 |
"[| Says Server B |
|
135 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}, |
|
136 |
Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
137 |
\<in> set evs; |
|
138 |
evs \<in> otway |] |
|
139 |
==> K \<notin> range shrK & (\<exists>i. NA = Nonce i) & (\<exists>j. NB = Nonce j)" |
|
140 |
apply (erule rev_mp) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
141 |
apply (erule otway.induct, auto) |
11251 | 142 |
done |
143 |
||
144 |
||
145 |
||
146 |
(**** |
|
147 |
The following is to prove theorems of the form |
|
148 |
||
149 |
Key K \<in> analz (insert (Key KAB) (knows Spy evs)) ==> |
|
150 |
Key K \<in> analz (knows Spy evs) |
|
151 |
||
152 |
A more general formula must be proved inductively. |
|
153 |
****) |
|
154 |
||
155 |
||
14238 | 156 |
text{* Session keys are not used to encrypt other session keys *} |
11251 | 157 |
|
14238 | 158 |
text{*The equality makes the induction hypothesis easier to apply*} |
11251 | 159 |
lemma analz_image_freshK [rule_format]: |
160 |
"evs \<in> otway ==> |
|
161 |
\<forall>K KK. KK <= -(range shrK) --> |
|
162 |
(Key K \<in> analz (Key`KK Un (knows Spy evs))) = |
|
163 |
(K \<in> KK | Key K \<in> analz (knows Spy evs))" |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
164 |
apply (erule otway.induct) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
165 |
apply (frule_tac [8] Says_Server_message_form) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
166 |
apply (drule_tac [7] OR4_analz_knows_Spy, analz_freshK, spy_analz, auto) |
11251 | 167 |
done |
168 |
||
169 |
lemma analz_insert_freshK: |
|
170 |
"[| evs \<in> otway; KAB \<notin> range shrK |] ==> |
|
11655 | 171 |
(Key K \<in> analz (insert (Key KAB) (knows Spy evs))) = |
11251 | 172 |
(K = KAB | Key K \<in> analz (knows Spy evs))" |
173 |
by (simp only: analz_image_freshK analz_image_freshK_simps) |
|
174 |
||
175 |
||
14238 | 176 |
text{*The Key K uniquely identifies the Server's message.*} |
11251 | 177 |
lemma unique_session_keys: |
178 |
"[| Says Server B |
|
179 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, K|}, |
|
180 |
Crypt (shrK B) {|NB, Agent A, Agent B, K|}|} |
|
181 |
\<in> set evs; |
|
182 |
Says Server B' |
|
183 |
{|Crypt (shrK A') {|NA', Agent A', Agent B', K|}, |
|
184 |
Crypt (shrK B') {|NB', Agent A', Agent B', K|}|} |
|
185 |
\<in> set evs; |
|
186 |
evs \<in> otway |] |
|
187 |
==> A=A' & B=B' & NA=NA' & NB=NB'" |
|
13507 | 188 |
apply (erule rev_mp, erule rev_mp, erule otway.induct, simp_all) |
14238 | 189 |
apply blast+ --{*OR3 and OR4*} |
11251 | 190 |
done |
191 |
||
192 |
||
14238 | 193 |
subsection{*Authenticity properties relating to NA*} |
11251 | 194 |
|
14238 | 195 |
text{*If the encrypted message appears then it originated with the Server!*} |
11251 | 196 |
lemma NA_Crypt_imp_Server_msg [rule_format]: |
197 |
"[| A \<notin> bad; A \<noteq> B; evs \<in> otway |] |
|
198 |
==> Crypt (shrK A) {|NA, Agent A, Agent B, Key K|} \<in> parts (knows Spy evs) |
|
199 |
--> (\<exists>NB. Says Server B |
|
200 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}, |
|
201 |
Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
202 |
\<in> set evs)" |
|
203 |
apply (erule otway.induct, force) |
|
204 |
apply (simp_all add: ex_disj_distrib) |
|
14238 | 205 |
apply blast+ --{*Fake, OR3*} |
11251 | 206 |
done |
207 |
||
208 |
||
14238 | 209 |
text{*Corollary: if A receives B's OR4 message then it originated with the |
210 |
Server. Freshness may be inferred from nonce NA.*} |
|
11251 | 211 |
lemma A_trusts_OR4: |
212 |
"[| Says B' A (Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}) \<in> set evs; |
|
213 |
A \<notin> bad; A \<noteq> B; evs \<in> otway |] |
|
214 |
==> \<exists>NB. Says Server B |
|
215 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}, |
|
216 |
Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
217 |
\<in> set evs" |
|
218 |
by (blast intro!: NA_Crypt_imp_Server_msg) |
|
219 |
||
220 |
||
14238 | 221 |
text{*Crucial secrecy property: Spy does not see the keys sent in msg OR3 |
11251 | 222 |
Does not in itself guarantee security: an attack could violate |
14238 | 223 |
the premises, e.g. by having @{term "A=Spy"}*} |
11251 | 224 |
lemma secrecy_lemma: |
225 |
"[| A \<notin> bad; B \<notin> bad; evs \<in> otway |] |
|
226 |
==> Says Server B |
|
227 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}, |
|
228 |
Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
229 |
\<in> set evs --> |
|
230 |
Notes Spy {|NA, NB, Key K|} \<notin> set evs --> |
|
231 |
Key K \<notin> analz (knows Spy evs)" |
|
232 |
apply (erule otway.induct, force) |
|
233 |
apply (frule_tac [7] Says_Server_message_form) |
|
234 |
apply (drule_tac [6] OR4_analz_knows_Spy) |
|
14238 | 235 |
apply (simp_all add: analz_insert_eq analz_insert_freshK pushes) |
236 |
apply spy_analz --{*Fake*} |
|
237 |
apply (blast dest: unique_session_keys)+ --{*OR3, OR4, Oops*} |
|
11251 | 238 |
done |
239 |
||
240 |
||
241 |
lemma Spy_not_see_encrypted_key: |
|
242 |
"[| Says Server B |
|
243 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}, |
|
244 |
Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
245 |
\<in> set evs; |
|
246 |
Notes Spy {|NA, NB, Key K|} \<notin> set evs; |
|
247 |
A \<notin> bad; B \<notin> bad; evs \<in> otway |] |
|
248 |
==> Key K \<notin> analz (knows Spy evs)" |
|
32366
b269b56b6a14
Demonstrations of sledgehammer in protocol proofs.
paulson
parents:
23746
diff
changeset
|
249 |
by (metis secrecy_lemma) |
11251 | 250 |
|
251 |
||
14238 | 252 |
text{*A's guarantee. The Oops premise quantifies over NB because A cannot know |
253 |
what it is.*} |
|
11251 | 254 |
lemma A_gets_good_key: |
255 |
"[| Says B' A (Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}) \<in> set evs; |
|
256 |
\<forall>NB. Notes Spy {|NA, NB, Key K|} \<notin> set evs; |
|
257 |
A \<notin> bad; B \<notin> bad; A \<noteq> B; evs \<in> otway |] |
|
258 |
==> Key K \<notin> analz (knows Spy evs)" |
|
32366
b269b56b6a14
Demonstrations of sledgehammer in protocol proofs.
paulson
parents:
23746
diff
changeset
|
259 |
by (metis A_trusts_OR4 secrecy_lemma) |
11251 | 260 |
|
261 |
||
262 |
||
14238 | 263 |
subsection{*Authenticity properties relating to NB*} |
11251 | 264 |
|
14238 | 265 |
text{*If the encrypted message appears then it originated with the Server!*} |
11251 | 266 |
lemma NB_Crypt_imp_Server_msg [rule_format]: |
267 |
"[| B \<notin> bad; A \<noteq> B; evs \<in> otway |] |
|
268 |
==> Crypt (shrK B) {|NB, Agent A, Agent B, Key K|} \<in> parts (knows Spy evs) |
|
269 |
--> (\<exists>NA. Says Server B |
|
270 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}, |
|
271 |
Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
272 |
\<in> set evs)" |
|
273 |
apply (erule otway.induct, force, simp_all add: ex_disj_distrib) |
|
14238 | 274 |
apply blast+ --{*Fake, OR3*} |
11251 | 275 |
done |
276 |
||
277 |
||
278 |
||
14238 | 279 |
text{*Guarantee for B: if it gets a well-formed certificate then the Server |
280 |
has sent the correct message in round 3.*} |
|
11251 | 281 |
lemma B_trusts_OR3: |
282 |
"[| Says S B {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
283 |
\<in> set evs; |
|
284 |
B \<notin> bad; A \<noteq> B; evs \<in> otway |] |
|
285 |
==> \<exists>NA. Says Server B |
|
286 |
{|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}, |
|
287 |
Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
288 |
\<in> set evs" |
|
289 |
by (blast intro!: NB_Crypt_imp_Server_msg) |
|
290 |
||
291 |
||
14238 | 292 |
text{*The obvious combination of @{text B_trusts_OR3} with |
293 |
@{text Spy_not_see_encrypted_key}*} |
|
11251 | 294 |
lemma B_gets_good_key: |
295 |
"[| Gets B {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} |
|
296 |
\<in> set evs; |
|
297 |
\<forall>NA. Notes Spy {|NA, NB, Key K|} \<notin> set evs; |
|
298 |
A \<notin> bad; B \<notin> bad; A \<noteq> B; evs \<in> otway |] |
|
299 |
==> Key K \<notin> analz (knows Spy evs)" |
|
300 |
by (blast dest: B_trusts_OR3 Spy_not_see_encrypted_key) |
|
2090 | 301 |
|
302 |
end |