author | paulson |
Tue, 08 Sep 1998 15:17:11 +0200 | |
changeset 5434 | 9b4bed3f394c |
parent 5359 | bd539b72d484 |
child 6335 | 7e4bffaa2a3e |
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 |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3481
diff
changeset
|
15 |
consts yahalom :: event list set |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3481
diff
changeset
|
16 |
inductive "yahalom" |
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*) |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3481
diff
changeset
|
19 |
Nil "[]: yahalom" |
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.*) |
5434
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5359
diff
changeset
|
24 |
Fake "[| evs: yahalom; X: synth (analz (spies evs)) |] |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3481
diff
changeset
|
25 |
==> Says Spy B X # evs : yahalom" |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
26 |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
27 |
(*Alice initiates a protocol run*) |
5434
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5359
diff
changeset
|
28 |
YM1 "[| evs1: yahalom; Nonce NA ~: used evs1 |] |
3659
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
29 |
==> Says A B {|Agent A, Nonce NA|} # evs1 : yahalom" |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
30 |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
31 |
(*Bob's response to Alice's message. Bob doesn't know who |
1995 | 32 |
the sender is, hence the A' in the sender field.*) |
5434
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5359
diff
changeset
|
33 |
YM2 "[| evs2: yahalom; Nonce NB ~: used evs2; |
3659
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
34 |
Says A' B {|Agent A, Nonce NA|} : set evs2 |] |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
35 |
==> Says B Server |
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
36 |
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|} |
3659
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
37 |
# evs2 : yahalom" |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
38 |
|
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
39 |
(*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
|
40 |
new session key to Alice, with a packet for forwarding to Bob.*) |
5434
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5359
diff
changeset
|
41 |
YM3 "[| evs3: yahalom; Key KAB ~: used evs3; |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
42 |
Says B' Server |
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2156
diff
changeset
|
43 |
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|} |
3659
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
44 |
: set evs3 |] |
1995 | 45 |
==> Says Server A |
3447
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
46 |
{|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|}, |
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
47 |
Crypt (shrK B) {|Agent A, Key KAB|}|} |
3659
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
48 |
# evs3 : yahalom" |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
49 |
|
1995 | 50 |
(*Alice receives the Server's (?) message, checks her Nonce, and |
3961 | 51 |
uses the new session key to send Bob his Nonce. The premise |
5434
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5359
diff
changeset
|
52 |
A ~= Server is needed to prove Says_Server_not_range.*) |
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5359
diff
changeset
|
53 |
YM4 "[| evs4: yahalom; A ~= Server; |
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|}, |
3659
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
55 |
X|} : set evs4; |
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
56 |
Says A B {|Agent A, Nonce NA|} : set evs4 |] |
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
57 |
==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 : yahalom" |
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.*) |
|
5359 | 62 |
Oops "[| evso: yahalom; |
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|}, |
3659
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
paulson
parents:
3519
diff
changeset
|
65 |
X|} : set evso |] |
4537
4e835bd9fada
Expressed most Oops rules using Notes instead of Says, and other tidying
paulson
parents:
3961
diff
changeset
|
66 |
==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso : yahalom" |
2110 | 67 |
|
3447
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
68 |
|
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
69 |
constdefs |
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
70 |
KeyWithNonce :: [key, nat, event list] => bool |
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
71 |
"KeyWithNonce K NB evs == |
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
72 |
EX A B na X. |
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
73 |
Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB|}, X|} |
3465 | 74 |
: set evs" |
3447
c7c8c0db05b9
Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents:
2516
diff
changeset
|
75 |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff
changeset
|
76 |
end |