author | paulson |
Thu, 26 Sep 1996 12:50:48 +0200 | |
changeset 2032 | 1bbf1bdcaf56 |
parent 1995 | c80e58e78d9c |
child 2110 | d01151e66cd4 |
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 |
|
2032 | 15 |
consts yahalom :: "agent set => event list set" |
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*) |
2032 | 29 |
YM1 "[| evs: yahalom lost; A ~= B |] |
30 |
==> Says A B {|Agent A, Nonce (newN evs)|} # 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.*) |
2032 | 34 |
YM2 "[| evs: yahalom lost; B ~= Server; |
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 |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
37 |
{|Agent B, |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
38 |
Crypt {|Agent A, Nonce NA, Nonce (newN evs)|} (shrK B)|} |
2032 | 39 |
# evs : yahalom lost" |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
40 |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
41 |
(*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
|
42 |
new session key to Alice, with a packet for forwarding to Bob.*) |
2032 | 43 |
YM3 "[| evs: yahalom lost; A ~= Server; |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
44 |
Says B' Server |
1995 | 45 |
{|Agent B, Crypt {|Agent A, Nonce NA, Nonce NB|} (shrK B)|} |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
46 |
: set_of_list evs |] |
1995 | 47 |
==> Says Server A |
48 |
{|Crypt {|Agent B, Key (newK evs), |
|
49 |
Nonce NA, Nonce NB|} (shrK A), |
|
50 |
Crypt {|Agent A, Key (newK evs)|} (shrK B)|} |
|
2032 | 51 |
# evs : yahalom lost" |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
52 |
|
1995 | 53 |
(*Alice receives the Server's (?) message, checks her Nonce, and |
54 |
uses the new session key to send Bob his Nonce.*) |
|
2032 | 55 |
YM4 "[| evs: yahalom lost; A ~= B; |
1995 | 56 |
Says S A {|Crypt {|Agent B, Key K, Nonce NA, Nonce NB|} (shrK A), |
57 |
X|} |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
58 |
: set_of_list evs; |
1995 | 59 |
Says A B {|Agent A, Nonce NA|} : set_of_list evs |] |
2032 | 60 |
==> Says A B {|X, Crypt (Nonce NB) K|} # evs : yahalom lost" |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
61 |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
62 |
end |