| author | paulson |
| Wed, 07 May 1997 16:29:06 +0200 | |
| changeset 3128 | d01d4c0c4b44 |
| parent 2516 | 4d68fbe6378b |
| child 3447 | c7c8c0db05b9 |
| permissions | -rw-r--r-- |
| 1995 | 1 |
(* Title: HOL/Auth/Yahalom |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
4 |
Copyright 1996 University of Cambridge |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
5 |
|
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
6 |
Inductive relation "yahalom" for the Yahalom protocol. |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
7 |
|
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
8 |
From page 257 of |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
9 |
Burrows, Abadi and Needham. A Logic of Authentication. |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
10 |
Proc. Royal Soc. 426 (1989) |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
11 |
*) |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
12 |
|
| 1995 | 13 |
Yahalom = Shared + |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
14 |
|
| 2378 | 15 |
consts yahalom :: agent set => event list set |
| 2032 | 16 |
inductive "yahalom lost" |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
17 |
intrs |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
18 |
(*Initial trace is empty*) |
| 2032 | 19 |
Nil "[]: yahalom lost" |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
20 |
|
| 2032 | 21 |
(*The spy MAY say anything he CAN say. We do not expect him to |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
22 |
invent new nonces here, but he can also use NS1. Common to |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
23 |
all similar protocols.*) |
| 2032 | 24 |
Fake "[| evs: yahalom lost; B ~= Spy; |
25 |
X: synth (analz (sees lost Spy evs)) |] |
|
26 |
==> Says Spy B X # evs : yahalom lost" |
|
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
27 |
|
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
28 |
(*Alice initiates a protocol run*) |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
29 |
YM1 "[| evs: yahalom lost; A ~= B; Nonce NA ~: used evs |] |
|
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
30 |
==> Says A B {|Agent A, Nonce NA|} # evs : yahalom lost"
|
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
31 |
|
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
32 |
(*Bob's response to Alice's message. Bob doesn't know who |
| 1995 | 33 |
the sender is, hence the A' in the sender field.*) |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
34 |
YM2 "[| evs: yahalom lost; B ~= Server; Nonce NB ~: used evs; |
| 1995 | 35 |
Says A' B {|Agent A, Nonce NA|} : set_of_list evs |]
|
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
36 |
==> Says B Server |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
37 |
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
|
|
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
38 |
# evs : yahalom lost" |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
39 |
|
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
40 |
(*The Server receives Bob's message. He responds by sending a |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
41 |
new session key to Alice, with a packet for forwarding to Bob.*) |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
42 |
YM3 "[| evs: yahalom lost; A ~= Server; Key KAB ~: used evs; |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
43 |
Says B' Server |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2156
diff
changeset
|
44 |
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
|
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
45 |
: set_of_list evs |] |
| 1995 | 46 |
==> Says Server A |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
47 |
{|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|},
|
|
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
48 |
Crypt (shrK B) {|Agent A, Key KAB|}|}
|
| 2032 | 49 |
# evs : yahalom lost" |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
50 |
|
| 1995 | 51 |
(*Alice receives the Server's (?) message, checks her Nonce, and |
52 |
uses the new session key to send Bob his Nonce.*) |
|
| 2125 | 53 |
YM4 "[| evs: yahalom lost; A ~= Server; A ~= B; |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2156
diff
changeset
|
54 |
Says S A {|Crypt (shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|},
|
| 2125 | 55 |
X|} : set_of_list evs; |
| 1995 | 56 |
Says A B {|Agent A, Nonce NA|} : set_of_list evs |]
|
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2156
diff
changeset
|
57 |
==> Says A B {|X, Crypt K (Nonce NB)|} # evs : yahalom lost"
|
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
58 |
|
| 2110 | 59 |
(*This message models possible leaks of session keys. The Nonces |
| 2156 | 60 |
identify the protocol run. Quoting Server here ensures they are |
61 |
correct.*) |
|
| 2125 | 62 |
Oops "[| evs: yahalom lost; A ~= Spy; |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2156
diff
changeset
|
63 |
Says Server A {|Crypt (shrK A)
|
|
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2156
diff
changeset
|
64 |
{|Agent B, Key K, Nonce NA, Nonce NB|},
|
| 2125 | 65 |
X|} : set_of_list evs |] |
| 2110 | 66 |
==> Says A Spy {|Nonce NA, Nonce NB, Key K|} # evs : yahalom lost"
|
67 |
||
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
68 |
end |