src/HOL/Auth/OtwayRees_Bad.thy
changeset 2451 ce85a2aafc7a
parent 2284 80ebd1a213fd
child 2516 4d68fbe6378b
equal deleted inserted replaced
2450:3ad2493fa0e0 2451:ce85a2aafc7a
    26     Fake "[| evs: otway;  B ~= Spy;  X: synth (analz (sees lost Spy evs)) |]
    26     Fake "[| evs: otway;  B ~= Spy;  X: synth (analz (sees lost Spy evs)) |]
    27           ==> Says Spy B X  # evs : otway"
    27           ==> Says Spy B X  # evs : otway"
    28 
    28 
    29          (*Alice initiates a protocol run*)
    29          (*Alice initiates a protocol run*)
    30     OR1  "[| evs: otway;  A ~= B;  B ~= Server |]
    30     OR1  "[| evs: otway;  A ~= B;  B ~= Server |]
    31           ==> Says A B {|Nonce (newN evs), Agent A, Agent B, 
    31           ==> Says A B {|Nonce (newN(length evs)), Agent A, Agent B, 
    32                          Crypt (shrK A) 
    32                          Crypt (shrK A) 
    33                                {|Nonce (newN evs), Agent A, Agent B|} |} 
    33                             {|Nonce (newN(length evs)), Agent A, Agent B|} |} 
    34                  # evs : otway"
    34                  # evs : otway"
    35 
    35 
    36          (*Bob's response to Alice's message.  Bob doesn't know who 
    36          (*Bob's response to Alice's message.  Bob doesn't know who 
    37 	   the sender is, hence the A' in the sender field.
    37 	   the sender is, hence the A' in the sender field.
    38            We modify the published protocol by NOT encrypting NB.*)
    38            We modify the published protocol by NOT encrypting NB.*)
    39     OR2  "[| evs: otway;  B ~= Server;
    39     OR2  "[| evs: otway;  B ~= Server;
    40              Says A' B {|Nonce NA, Agent A, Agent B, X|} : set_of_list evs |]
    40              Says A' B {|Nonce NA, Agent A, Agent B, X|} : set_of_list evs |]
    41           ==> Says B Server 
    41           ==> Says B Server 
    42                   {|Nonce NA, Agent A, Agent B, X, Nonce (newN evs), 
    42                   {|Nonce NA, Agent A, Agent B, X, Nonce (newN(length evs)), 
    43                     Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
    43                     Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
    44                  # evs : otway"
    44                  # evs : otway"
    45 
    45 
    46          (*The Server receives Bob's message and checks that the three NAs
    46          (*The Server receives Bob's message and checks that the three NAs
    47            match.  Then he sends a new session key to Bob with a packet for
    47            match.  Then he sends a new session key to Bob with a packet for
    53                     Nonce NB, 
    53                     Nonce NB, 
    54                     Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
    54                     Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
    55                : set_of_list evs |]
    55                : set_of_list evs |]
    56           ==> Says Server B 
    56           ==> Says Server B 
    57                   {|Nonce NA, 
    57                   {|Nonce NA, 
    58                     Crypt (shrK A) {|Nonce NA, Key (newK evs)|},
    58                     Crypt (shrK A) {|Nonce NA, Key (newK(length evs))|},
    59                     Crypt (shrK B) {|Nonce NB, Key (newK evs)|}|}
    59                     Crypt (shrK B) {|Nonce NB, Key (newK(length evs))|}|}
    60                  # evs : otway"
    60                  # evs : otway"
    61 
    61 
    62          (*Bob receives the Server's (?) message and compares the Nonces with
    62          (*Bob receives the Server's (?) message and compares the Nonces with
    63 	   those in the message he previously sent the Server.*)
    63 	   those in the message he previously sent the Server.*)
    64     OR4  "[| evs: otway;  A ~= B;
    64     OR4  "[| evs: otway;  A ~= B;