src/HOL/Auth/Yahalom_Bad.thy
author nipkow
Tue, 17 Oct 2000 08:00:46 +0200
changeset 10228 e653cb933293
parent 6400 1f495d4d922b
child 11185 1b737b4c2108
permissions -rw-r--r--
added intermediate value thms
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6400
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Yahalom
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     5
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     6
Inductive relation "yahalom" for the Yahalom protocol.
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     7
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     8
Example of why Oops is necessary.  This protocol can be attacked because it
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
     9
doesn't keep NB secret, but without Oops it can be "verified" anyway.
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    10
*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    11
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    12
Yahalom_Bad = Shared + 
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    13
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    14
consts  yahalom   :: event list set
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    15
inductive "yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    16
  intrs 
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    17
         (*Initial trace is empty*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    18
    Nil  "[]: yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    19
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    20
         (*The spy MAY say anything he CAN say.  We do not expect him to
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    21
           invent new nonces here, but he can also use NS1.  Common to
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    22
           all similar protocols.*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    23
    Fake "[| evs: yahalom;  X: synth (analz (knows Spy evs)) |]
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    24
          ==> Says Spy B X  # evs : yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    25
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    26
         (*A message that has been sent can be received by the
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    27
           intended recipient.*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    28
    Reception "[| evsr: yahalom;  Says A B X : set evsr |]
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    29
               ==> Gets B X # evsr : yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    30
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    31
         (*Alice initiates a protocol run*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    32
    YM1  "[| evs1: yahalom;  Nonce NA ~: used evs1 |]
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    33
          ==> Says A B {|Agent A, Nonce NA|} # evs1 : yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    34
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    35
         (*Bob's response to Alice's message.*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    36
    YM2  "[| evs2: yahalom;  Nonce NB ~: used evs2;
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    37
             Gets B {|Agent A, Nonce NA|} : set evs2 |]
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    38
          ==> Says B Server 
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    39
                  {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    40
                # evs2 : yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    41
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    42
         (*The Server receives Bob's message.  He responds by sending a
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    43
            new session key to Alice, with a packet for forwarding to Bob.*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    44
    YM3  "[| evs3: yahalom;  Key KAB ~: used evs3;
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    45
             Gets Server 
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    46
                  {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    47
               : set evs3 |]
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    48
          ==> Says Server A
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    49
                   {|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|},
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    50
                     Crypt (shrK B) {|Agent A, Key KAB|}|}
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    51
                # evs3 : yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    52
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    53
         (*Alice receives the Server's (?) message, checks her Nonce, and
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    54
           uses the new session key to send Bob his Nonce.  The premise
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    55
           A ~= Server is needed to prove Says_Server_not_range.*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    56
    YM4  "[| evs4: yahalom;  A ~= Server;
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    57
             Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    58
                : set evs4;
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    59
             Says A B {|Agent A, Nonce NA|} : set evs4 |]
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    60
          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 : yahalom"
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    61
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    62
         (*This message models possible leaks of session keys.  The Nonces
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    63
           identify the protocol run.  Quoting Server here ensures they are
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    64
           correct.*)
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    65
1f495d4d922b added new theory Yahalom_Bad
paulson
parents:
diff changeset
    66
end