src/HOL/Auth/OtwayRees.thy
changeset 3659 eddedfe2f3f8
parent 3519 ab0a9fbed4c0
child 3683 aafe719dff14
equal deleted inserted replaced
3658:f87dd7b68d8c 3659:eddedfe2f3f8
    26     Fake "[| evs: otway;  B ~= Spy;  
    26     Fake "[| evs: otway;  B ~= Spy;  
    27              X: synth (analz (sees Spy evs)) |]
    27              X: synth (analz (sees Spy evs)) |]
    28           ==> Says Spy B X  # evs : otway"
    28           ==> Says Spy B X  # evs : otway"
    29 
    29 
    30          (*Alice initiates a protocol run*)
    30          (*Alice initiates a protocol run*)
    31     OR1  "[| evs: otway;  A ~= B;  B ~= Server;  Nonce NA ~: used evs |]
    31     OR1  "[| evs1: otway;  A ~= B;  B ~= Server;  Nonce NA ~: used evs1 |]
    32           ==> Says A B {|Nonce NA, Agent A, Agent B, 
    32           ==> Says A B {|Nonce NA, Agent A, Agent B, 
    33                          Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |} 
    33                          Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |} 
    34                  # evs : otway"
    34                  # evs1 : 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            Note that NB is encrypted.*)
    38            Note that NB is encrypted.*)
    39     OR2  "[| evs: otway;  B ~= Server;  Nonce NB ~: used evs;
    39     OR2  "[| evs2: otway;  B ~= Server;  Nonce NB ~: used evs2;
    40              Says A' B {|Nonce NA, Agent A, Agent B, X|} : set evs |]
    40              Says A' B {|Nonce NA, Agent A, Agent B, X|} : set evs2 |]
    41           ==> Says B Server 
    41           ==> Says B Server 
    42                   {|Nonce NA, Agent A, Agent B, X, 
    42                   {|Nonce NA, Agent A, Agent B, X, 
    43                     Crypt (shrK B)
    43                     Crypt (shrK B)
    44                       {|Nonce NA, Nonce NB, Agent A, Agent B|}|}
    44                       {|Nonce NA, Nonce NB, Agent A, Agent B|}|}
    45                  # evs : otway"
    45                  # evs2 : otway"
    46 
    46 
    47          (*The Server receives Bob's message and checks that the three NAs
    47          (*The Server receives Bob's message and checks that the three NAs
    48            match.  Then he sends a new session key to Bob with a packet for
    48            match.  Then he sends a new session key to Bob with a packet for
    49            forwarding to Alice.*)
    49            forwarding to Alice.*)
    50     OR3  "[| evs: otway;  B ~= Server;  Key KAB ~: used evs;
    50     OR3  "[| evs3: otway;  B ~= Server;  Key KAB ~: used evs3;
    51              Says B' Server 
    51              Says B' Server 
    52                   {|Nonce NA, Agent A, Agent B, 
    52                   {|Nonce NA, Agent A, Agent B, 
    53                     Crypt (shrK A) {|Nonce NA, Agent A, Agent B|}, 
    53                     Crypt (shrK A) {|Nonce NA, Agent A, Agent B|}, 
    54                     Crypt (shrK B) {|Nonce NA, Nonce NB, Agent A, Agent B|}|}
    54                     Crypt (shrK B) {|Nonce NA, Nonce NB, Agent A, Agent B|}|}
    55                : set evs |]
    55                : set evs3 |]
    56           ==> Says Server B 
    56           ==> Says Server B 
    57                   {|Nonce NA, 
    57                   {|Nonce NA, 
    58                     Crypt (shrK A) {|Nonce NA, Key KAB|},
    58                     Crypt (shrK A) {|Nonce NA, Key KAB|},
    59                     Crypt (shrK B) {|Nonce NB, Key KAB|}|}
    59                     Crypt (shrK B) {|Nonce NB, Key KAB|}|}
    60                  # evs : otway"
    60                  # evs3 : 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  "[| evs4: otway;  A ~= B;  
    65              Says B Server {|Nonce NA, Agent A, Agent B, X', 
    65              Says B Server {|Nonce NA, Agent A, Agent B, X', 
    66                              Crypt (shrK B)
    66                              Crypt (shrK B)
    67                                    {|Nonce NA, Nonce NB, Agent A, Agent B|}|}
    67                                    {|Nonce NA, Nonce NB, Agent A, Agent B|}|}
    68                : set evs;
    68                : set evs4;
    69              Says S' B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
    69              Says S' B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
    70                : set evs |]
    70                : set evs4 |]
    71           ==> Says B A {|Nonce NA, X|} # evs : otway"
    71           ==> Says B A {|Nonce NA, X|} # evs4 : otway"
    72 
    72 
    73          (*This message models possible leaks of session keys.  The nonces
    73          (*This message models possible leaks of session keys.  The nonces
    74            identify the protocol run.*)
    74            identify the protocol run.*)
    75     Oops "[| evs: otway;  B ~= Spy;
    75     Oops "[| evso: otway;  B ~= Spy;
    76              Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
    76              Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
    77                : set evs |]
    77                : set evso |]
    78           ==> Says B Spy {|Nonce NA, Nonce NB, Key K|} # evs : otway"
    78           ==> Says B Spy {|Nonce NA, Nonce NB, Key K|} # evso : otway"
    79 
    79 
    80 end
    80 end