--- a/src/HOL/Auth/Yahalom.thy Tue Sep 08 14:54:21 1998 +0200
+++ b/src/HOL/Auth/Yahalom.thy Tue Sep 08 15:17:11 1998 +0200
@@ -21,17 +21,16 @@
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake "[| evs: yahalom; B ~= Spy;
- X: synth (analz (spies evs)) |]
+ Fake "[| evs: yahalom; X: synth (analz (spies evs)) |]
==> Says Spy B X # evs : yahalom"
(*Alice initiates a protocol run*)
- YM1 "[| evs1: yahalom; A ~= B; Nonce NA ~: used evs1 |]
+ YM1 "[| evs1: yahalom; Nonce NA ~: used evs1 |]
==> Says A B {|Agent A, Nonce NA|} # evs1 : yahalom"
(*Bob's response to Alice's message. Bob doesn't know who
the sender is, hence the A' in the sender field.*)
- YM2 "[| evs2: yahalom; B ~= Server; Nonce NB ~: used evs2;
+ YM2 "[| evs2: yahalom; Nonce NB ~: used evs2;
Says A' B {|Agent A, Nonce NA|} : set evs2 |]
==> Says B Server
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
@@ -39,7 +38,7 @@
(*The Server receives Bob's message. He responds by sending a
new session key to Alice, with a packet for forwarding to Bob.*)
- YM3 "[| evs3: yahalom; A ~= Server; Key KAB ~: used evs3;
+ YM3 "[| evs3: yahalom; Key KAB ~: used evs3;
Says B' Server
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
: set evs3 |]
@@ -50,8 +49,8 @@
(*Alice receives the Server's (?) message, checks her Nonce, and
uses the new session key to send Bob his Nonce. The premise
- A ~= Server is needed to prove Says_Server_message_form.*)
- YM4 "[| evs4: yahalom; A ~= Server;
+ A ~= Server is needed to prove Says_Server_not_range.*)
+ YM4 "[| evs4: yahalom; A ~= Server;
Says S A {|Crypt (shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|},
X|} : set evs4;
Says A B {|Agent A, Nonce NA|} : set evs4 |]