src/HOL/Auth/Yahalom2.thy
changeset 2111 81c8d46edfa3
child 2155 dc85854810eb
equal deleted inserted replaced
2110:d01151e66cd4 2111:81c8d46edfa3
       
     1 (*  Title:      HOL/Auth/Yahalom
       
     2     ID:         $Id$
       
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1996  University of Cambridge
       
     5 
       
     6 Inductive relation "yahalom" for the Yahalom protocol, Variant 2.
       
     7 
       
     8 This version trades encryption of NB for additional explicitness in YM3.
       
     9 
       
    10 From page 259 of
       
    11   Burrows, Abadi and Needham.  A Logic of Authentication.
       
    12   Proc. Royal Soc. 426 (1989)
       
    13 *)
       
    14 
       
    15 Yahalom2 = Shared + 
       
    16 
       
    17 consts  yahalom   :: "agent set => event list set"
       
    18 inductive "yahalom lost"
       
    19   intrs 
       
    20          (*Initial trace is empty*)
       
    21     Nil  "[]: yahalom lost"
       
    22 
       
    23          (*The spy MAY say anything he CAN say.  We do not expect him to
       
    24            invent new nonces here, but he can also use NS1.  Common to
       
    25            all similar protocols.*)
       
    26     Fake "[| evs: yahalom lost;  B ~= Spy;  
       
    27              X: synth (analz (sees lost Spy evs)) |]
       
    28           ==> Says Spy B X  # evs : yahalom lost"
       
    29 
       
    30          (*Alice initiates a protocol run*)
       
    31     YM1  "[| evs: yahalom lost;  A ~= B |]
       
    32           ==> Says A B {|Agent A, Nonce (newN evs)|} # evs : yahalom lost"
       
    33 
       
    34          (*Bob's response to Alice's message.  Bob doesn't know who 
       
    35 	   the sender is, hence the A' in the sender field.*)
       
    36     YM2  "[| evs: yahalom lost;  B ~= Server;
       
    37              Says A' B {|Agent A, Nonce NA|} : set_of_list evs |]
       
    38           ==> Says B Server 
       
    39                   {|Agent B, Nonce (newN evs), 
       
    40                     Crypt {|Agent A, Nonce NA|} (shrK B)|}
       
    41                  # evs : yahalom lost"
       
    42 
       
    43          (*The Server receives Bob's message.  He responds by sending a
       
    44             new session key to Alice, with a packet for forwarding to Bob.*)
       
    45     YM3  "[| evs: yahalom lost;  A ~= Server;
       
    46              Says B' Server 
       
    47                   {|Agent B, Nonce NB, Crypt {|Agent A, Nonce NA|} (shrK B)|}
       
    48                : set_of_list evs |]
       
    49           ==> Says Server A
       
    50                {|Nonce NB, 
       
    51                  Crypt {|Agent B, Key (newK evs), Nonce NA|} (shrK A),
       
    52                  Crypt {|Agent A, Key (newK evs), Nonce NB, Nonce NB|} (shrK B)|}
       
    53                  # evs : yahalom lost"
       
    54 
       
    55          (*Alice receives the Server's (?) message, checks her Nonce, and
       
    56            uses the new session key to send Bob his Nonce.*)
       
    57     YM4  "[| evs: yahalom lost;  A ~= B;  
       
    58              Says S A {|Nonce NB, Crypt {|Agent B, Key K, Nonce NA|} (shrK A),
       
    59                         X|}
       
    60                : set_of_list evs;
       
    61              Says A B {|Agent A, Nonce NA|} : set_of_list evs |]
       
    62           ==> Says A B {|X, Crypt (Nonce NB) K|} # evs : yahalom lost"
       
    63 
       
    64          (*This message models possible leaks of session keys.  The Nonce NA
       
    65            identifies the protocol run.  We can't be sure about NB.*)
       
    66     Revl "[| evs: yahalom lost;  A ~= Spy;
       
    67              Says S A {|Nonce NB, Crypt {|Agent B, Key K, Nonce NA|} (shrK A),
       
    68                         X|}
       
    69                : set_of_list evs |]
       
    70           ==> Says A Spy {|Nonce NA, Key K|} # evs : yahalom lost"
       
    71 
       
    72 end