diff -r 6ff9bd353121 -r 4d68fbe6378b src/HOL/Auth/Yahalom2.thy --- a/src/HOL/Auth/Yahalom2.thy Fri Jan 17 11:50:09 1997 +0100 +++ b/src/HOL/Auth/Yahalom2.thy Fri Jan 17 12:49:31 1997 +0100 @@ -30,26 +30,26 @@ ==> Says Spy B X # evs : yahalom lost" (*Alice initiates a protocol run*) - YM1 "[| evs: yahalom lost; A ~= B |] - ==> Says A B {|Agent A, Nonce (newN(length evs))|} # evs : yahalom lost" + YM1 "[| evs: yahalom lost; A ~= B; Nonce NA ~: used evs |] + ==> Says A B {|Agent A, Nonce NA|} # 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; + YM2 "[| evs: yahalom lost; B ~= Server; Nonce NB ~: used evs; Says A' B {|Agent A, Nonce NA|} : set_of_list evs |] - ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce (newN(length evs))|} - # evs : yahalom lost" + ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} # evs + : yahalom lost" (*The Server receives Bob's message. He responds by sending a new session key to Alice, with a packet for forwarding to Bob. Fields are reversed in the 2nd packet to prevent attacks.*) - YM3 "[| evs: yahalom lost; A ~= B; A ~= Server; + YM3 "[| evs: yahalom lost; A ~= B; A ~= Server; Key KAB ~: used evs; Says B' Server {|Agent A, Agent B, Nonce NA, Nonce NB|} : set_of_list evs |] ==> Says Server A {|Nonce NB, - Crypt (shrK A) {|Agent B, Key (newK(length evs)), Nonce NA|}, - Crypt (shrK B) {|Nonce NB, Key (newK(length evs)), Agent A|}|} + Crypt (shrK A) {|Agent B, Key KAB, Nonce NA|}, + Crypt (shrK B) {|Nonce NB, Key KAB, Agent A|}|} # evs : yahalom lost" (*Alice receives the Server's (?) message, checks her Nonce, and