diff -r 3ad2493fa0e0 -r ce85a2aafc7a src/HOL/Auth/Yahalom2.thy --- a/src/HOL/Auth/Yahalom2.thy Thu Dec 19 11:54:19 1996 +0100 +++ b/src/HOL/Auth/Yahalom2.thy Thu Dec 19 11:58:39 1996 +0100 @@ -31,13 +31,13 @@ (*Alice initiates a protocol run*) YM1 "[| evs: yahalom lost; A ~= B |] - ==> Says A B {|Agent A, Nonce (newN evs)|} # evs : yahalom lost" + ==> Says A B {|Agent A, Nonce (newN(length evs))|} # evs : yahalom lost" (*Bob's response to Alice's message. Bob doesn't know who the sender is, hence the A' in the sender field.*) YM2 "[| evs: yahalom lost; B ~= Server; Says A' B {|Agent A, Nonce NA|} : set_of_list evs |] - ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce (newN evs)|} + ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce (newN(length evs))|} # evs : yahalom lost" (*The Server receives Bob's message. He responds by sending a @@ -48,8 +48,8 @@ : set_of_list evs |] ==> Says Server A {|Nonce NB, - Crypt (shrK A) {|Agent B, Key (newK evs), Nonce NA|}, - Crypt (shrK B) {|Nonce NB, Key (newK evs), Agent A|}|} + Crypt (shrK A) {|Agent B, Key (newK(length evs)), Nonce NA|}, + Crypt (shrK B) {|Nonce NB, Key (newK(length evs)), Agent A|}|} # evs : yahalom lost" (*Alice receives the Server's (?) message, checks her Nonce, and