6400
|
1 |
(* Title: HOL/Auth/Yahalom
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1996 University of Cambridge
|
|
5 |
|
|
6 |
Inductive relation "yahalom" for the Yahalom protocol.
|
|
7 |
|
|
8 |
Example of why Oops is necessary. This protocol can be attacked because it
|
|
9 |
doesn't keep NB secret, but without Oops it can be "verified" anyway.
|
|
10 |
*)
|
|
11 |
|
|
12 |
Yahalom_Bad = Shared +
|
|
13 |
|
|
14 |
consts yahalom :: event list set
|
|
15 |
inductive "yahalom"
|
|
16 |
intrs
|
|
17 |
(*Initial trace is empty*)
|
|
18 |
Nil "[]: yahalom"
|
|
19 |
|
|
20 |
(*The spy MAY say anything he CAN say. We do not expect him to
|
|
21 |
invent new nonces here, but he can also use NS1. Common to
|
|
22 |
all similar protocols.*)
|
|
23 |
Fake "[| evs: yahalom; X: synth (analz (knows Spy evs)) |]
|
|
24 |
==> Says Spy B X # evs : yahalom"
|
|
25 |
|
|
26 |
(*A message that has been sent can be received by the
|
|
27 |
intended recipient.*)
|
|
28 |
Reception "[| evsr: yahalom; Says A B X : set evsr |]
|
|
29 |
==> Gets B X # evsr : yahalom"
|
|
30 |
|
|
31 |
(*Alice initiates a protocol run*)
|
|
32 |
YM1 "[| evs1: yahalom; Nonce NA ~: used evs1 |]
|
|
33 |
==> Says A B {|Agent A, Nonce NA|} # evs1 : yahalom"
|
|
34 |
|
|
35 |
(*Bob's response to Alice's message.*)
|
|
36 |
YM2 "[| evs2: yahalom; Nonce NB ~: used evs2;
|
|
37 |
Gets B {|Agent A, Nonce NA|} : set evs2 |]
|
|
38 |
==> Says B Server
|
|
39 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
|
|
40 |
# evs2 : yahalom"
|
|
41 |
|
|
42 |
(*The Server receives Bob's message. He responds by sending a
|
|
43 |
new session key to Alice, with a packet for forwarding to Bob.*)
|
|
44 |
YM3 "[| evs3: yahalom; Key KAB ~: used evs3;
|
|
45 |
Gets Server
|
|
46 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
|
|
47 |
: set evs3 |]
|
|
48 |
==> Says Server A
|
|
49 |
{|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|},
|
|
50 |
Crypt (shrK B) {|Agent A, Key KAB|}|}
|
|
51 |
# evs3 : yahalom"
|
|
52 |
|
|
53 |
(*Alice receives the Server's (?) message, checks her Nonce, and
|
|
54 |
uses the new session key to send Bob his Nonce. The premise
|
|
55 |
A ~= Server is needed to prove Says_Server_not_range.*)
|
|
56 |
YM4 "[| evs4: yahalom; A ~= Server;
|
|
57 |
Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
|
|
58 |
: set evs4;
|
|
59 |
Says A B {|Agent A, Nonce NA|} : set evs4 |]
|
|
60 |
==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 : yahalom"
|
|
61 |
|
|
62 |
(*This message models possible leaks of session keys. The Nonces
|
|
63 |
identify the protocol run. Quoting Server here ensures they are
|
|
64 |
correct.*)
|
|
65 |
|
|
66 |
end
|