src/HOL/Auth/Yahalom.thy
author paulson
Fri, 11 Jul 1997 13:26:15 +0200
changeset 3512 9dcb4daa15e8
parent 3481 256f38c01b98
child 3519 ab0a9fbed4c0
permissions -rw-r--r--
Moving common declarations and proofs from theories "Shared" and "Public" to "Event". NB the original "Event" theory was later renamed "Shared". Addition of the Notes constructor to datatype "event".
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1995
c80e58e78d9c Addition of Yahalom protocol
paulson
parents: 1985
diff changeset
     1
(*  Title:      HOL/Auth/Yahalom
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     2
    ID:         $Id$
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     5
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     6
Inductive relation "yahalom" for the Yahalom protocol.
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     7
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     8
From page 257 of
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
     9
  Burrows, Abadi and Needham.  A Logic of Authentication.
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    10
  Proc. Royal Soc. 426 (1989)
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    11
*)
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    12
1995
c80e58e78d9c Addition of Yahalom protocol
paulson
parents: 1985
diff changeset
    13
Yahalom = Shared + 
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    14
2378
fc103154ad8f Removed needless quotation marks
paulson
parents: 2284
diff changeset
    15
consts  yahalom   :: agent set => event list set
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 1995
diff changeset
    16
inductive "yahalom lost"
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    17
  intrs 
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    18
         (*Initial trace is empty*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 1995
diff changeset
    19
    Nil  "[]: yahalom lost"
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    20
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 1995
diff changeset
    21
         (*The spy MAY say anything he CAN say.  We do not expect him to
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    22
           invent new nonces here, but he can also use NS1.  Common to
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    23
           all similar protocols.*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 1995
diff changeset
    24
    Fake "[| evs: yahalom lost;  B ~= Spy;  
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 1995
diff changeset
    25
             X: synth (analz (sees lost Spy evs)) |]
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 1995
diff changeset
    26
          ==> Says Spy B X  # evs : yahalom lost"
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    27
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    28
         (*Alice initiates a protocol run*)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    29
    YM1  "[| evs: yahalom lost;  A ~= B;  Nonce NA ~: used evs |]
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    30
          ==> Says A B {|Agent A, Nonce NA|} # evs : yahalom lost"
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    31
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    32
         (*Bob's response to Alice's message.  Bob doesn't know who 
1995
c80e58e78d9c Addition of Yahalom protocol
paulson
parents: 1985
diff changeset
    33
	   the sender is, hence the A' in the sender field.*)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    34
    YM2  "[| evs: yahalom lost;  B ~= Server;  Nonce NB ~: used evs;
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3447
diff changeset
    35
             Says A' B {|Agent A, Nonce NA|} : set evs |]
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    36
          ==> Says B Server 
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    37
                  {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    38
                # evs : yahalom lost"
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    39
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    40
         (*The Server receives Bob's message.  He responds by sending a
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    41
            new session key to Alice, with a packet for forwarding to Bob.*)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    42
    YM3  "[| evs: yahalom lost;  A ~= Server;  Key KAB ~: used evs;
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    43
             Says B' Server 
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2156
diff changeset
    44
                  {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3447
diff changeset
    45
               : set evs |]
1995
c80e58e78d9c Addition of Yahalom protocol
paulson
parents: 1985
diff changeset
    46
          ==> Says Server A
3447
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    47
                   {|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|},
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    48
                     Crypt (shrK B) {|Agent A, Key KAB|}|}
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    49
                # evs : yahalom lost"
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    50
1995
c80e58e78d9c Addition of Yahalom protocol
paulson
parents: 1985
diff changeset
    51
         (*Alice receives the Server's (?) message, checks her Nonce, and
c80e58e78d9c Addition of Yahalom protocol
paulson
parents: 1985
diff changeset
    52
           uses the new session key to send Bob his Nonce.*)
3481
256f38c01b98 Deleted a redundant A~=B in rules that refer to a previous event
paulson
parents: 3465
diff changeset
    53
    YM4  "[| evs: yahalom lost;  A ~= Server;  
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2156
diff changeset
    54
             Says S A {|Crypt (shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|},
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3447
diff changeset
    55
                        X|}  : set evs;
e85c24717cad set_of_list -> set
nipkow
parents: 3447
diff changeset
    56
             Says A B {|Agent A, Nonce NA|} : set evs |]
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2156
diff changeset
    57
          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs : yahalom lost"
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    58
2110
d01151e66cd4 Addition of Reveal message
paulson
parents: 2032
diff changeset
    59
         (*This message models possible leaks of session keys.  The Nonces
2156
9c361df93bd5 Minor changes to comments
paulson
parents: 2125
diff changeset
    60
           identify the protocol run.  Quoting Server here ensures they are
9c361df93bd5 Minor changes to comments
paulson
parents: 2125
diff changeset
    61
           correct.*)
2125
92a08ee6a9cb New Oops message, with Server as source to ensure
paulson
parents: 2110
diff changeset
    62
    Oops "[| evs: yahalom lost;  A ~= Spy;
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2156
diff changeset
    63
             Says Server A {|Crypt (shrK A)
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2156
diff changeset
    64
                                   {|Agent B, Key K, Nonce NA, Nonce NB|},
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3447
diff changeset
    65
                             X|}  : set evs |]
2110
d01151e66cd4 Addition of Reveal message
paulson
parents: 2032
diff changeset
    66
          ==> Says A Spy {|Nonce NA, Nonce NB, Key K|} # evs : yahalom lost"
d01151e66cd4 Addition of Reveal message
paulson
parents: 2032
diff changeset
    67
3447
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    68
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    69
constdefs 
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    70
  KeyWithNonce :: [key, nat, event list] => bool
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    71
  "KeyWithNonce K NB evs ==
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    72
     EX A B na X. 
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    73
       Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB|}, X|} 
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3447
diff changeset
    74
         : set evs"
3447
c7c8c0db05b9 Defines KeyWithNonce, which is used to prove the secrecy of NB
paulson
parents: 2516
diff changeset
    75
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
diff changeset
    76
end