author | paulson |
Mon, 14 Jul 1997 12:47:21 +0200 | |
changeset 3519 | ab0a9fbed4c0 |
parent 3466 | 30791e5a69c4 |
child 3659 | eddedfe2f3f8 |
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 |
|
5 |
||
6 |
Inductive relation "otway" for the Otway-Rees protocol. |
|
7 |
||
8 |
Simplified version with minimal encryption but explicit messages |
|
9 |
||
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
10 |
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
|
11 |
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
|
12 |
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
|
13 |
properties. |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
14 |
|
2090 | 15 |
From page 11 of |
16 |
Abadi and Needham. Prudent Engineering Practice for Cryptographic Protocols. |
|
17 |
IEEE Trans. SE 22 (1), 1996 |
|
18 |
*) |
|
19 |
||
20 |
OtwayRees_AN = Shared + |
|
21 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
22 |
consts otway :: event list set |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
23 |
inductive "otway" |
2090 | 24 |
intrs |
25 |
(*Initial trace is empty*) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
26 |
Nil "[]: otway" |
2090 | 27 |
|
28 |
(*The spy MAY say anything he CAN say. We do not expect him to |
|
29 |
invent new nonces here, but he can also use NS1. Common to |
|
30 |
all similar protocols.*) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
31 |
Fake "[| evs: otway; B ~= Spy; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
32 |
X: synth (analz (sees Spy evs)) |] |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
33 |
==> Says Spy B X # evs : otway" |
2090 | 34 |
|
35 |
(*Alice initiates a protocol run*) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
36 |
OR1 "[| evs: otway; A ~= B; B ~= Server |] |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
37 |
==> Says A B {|Agent A, Agent B, Nonce NA|} # evs : otway" |
2090 | 38 |
|
39 |
(*Bob's response to Alice's message. Bob doesn't know who |
|
40 |
the sender is, hence the A' in the sender field.*) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
41 |
OR2 "[| evs: otway; B ~= Server; |
3465 | 42 |
Says A' B {|Agent A, Agent B, Nonce NA|} : set evs |] |
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
43 |
==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
44 |
# evs : otway" |
2090 | 45 |
|
46 |
(*The Server receives Bob's message. Then he sends a new |
|
47 |
session key to Bob with a packet for forwarding to Alice.*) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
48 |
OR3 "[| evs: otway; B ~= Server; A ~= B; Key KAB ~: used evs; |
2090 | 49 |
Says B' Server {|Agent A, Agent B, Nonce NA, Nonce NB|} |
3465 | 50 |
: set evs |] |
2090 | 51 |
==> Says Server B |
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
52 |
{|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
|
53 |
Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key KAB|}|} |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
54 |
# evs : otway" |
2090 | 55 |
|
56 |
(*Bob receives the Server's (?) message and compares the Nonces with |
|
57 |
those in the message he previously sent the Server.*) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
58 |
OR4 "[| evs: otway; A ~= B; |
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
59 |
Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} : set evs; |
2837
dee1c7f1f576
Trivial renamings (for consistency with CSFW papers)
paulson
parents:
2516
diff
changeset
|
60 |
Says S' B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|} |
3465 | 61 |
: set evs |] |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
62 |
==> Says B A X # evs : otway" |
2090 | 63 |
|
2131 | 64 |
(*This message models possible leaks of session keys. The nonces |
65 |
identify the protocol run. B is not assumed to know shrK A.*) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
66 |
Oops "[| evs: otway; B ~= Spy; |
2131 | 67 |
Says Server B |
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2131
diff
changeset
|
68 |
{|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|}, |
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2131
diff
changeset
|
69 |
Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|} |
3465 | 70 |
: set evs |] |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3466
diff
changeset
|
71 |
==> Says B Spy {|Nonce NA, Nonce NB, Key K|} # evs : otway" |
2090 | 72 |
|
73 |
end |