src/HOL/Auth/OtwayRees_Bad.thy
author paulson
Fri, 11 Jul 1997 13:26:15 +0200
changeset 3512 9dcb4daa15e8
parent 3465 e85c24717cad
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:
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/OtwayRees_Bad
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     2
    ID:         $Id$
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     5
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     6
Inductive relation "otway" for the Otway-Rees protocol.
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     7
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     8
The FAULTY version omitting encryption of Nonce NB, as suggested on page 247 of
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
     9
  Burrows, Abadi and Needham.  A Logic of Authentication.
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    10
  Proc. Royal Soc. 426 (1989)
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    11
*)
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    12
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    13
OtwayRees_Bad = Shared + 
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    14
2052
d9f7f4b2613e Working again with new theory Shared
paulson
parents: 2032
diff changeset
    15
consts  lost    :: agent set        (*No need for it to be a variable*)
d9f7f4b2613e Working again with new theory Shared
paulson
parents: 2032
diff changeset
    16
	otway   :: event list set
d9f7f4b2613e Working again with new theory Shared
paulson
parents: 2032
diff changeset
    17
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    18
inductive otway
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    19
  intrs 
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    20
         (*Initial trace is empty*)
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    21
    Nil  "[]: otway"
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    22
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2002
diff changeset
    23
         (*The spy MAY say anything he CAN say.  We do not expect him to
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    24
           invent new nonces here, but he can also use NS1.  Common to
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    25
           all similar protocols.*)
2052
d9f7f4b2613e Working again with new theory Shared
paulson
parents: 2032
diff changeset
    26
    Fake "[| evs: otway;  B ~= Spy;  X: synth (analz (sees lost Spy evs)) |]
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2002
diff changeset
    27
          ==> Says Spy B X  # evs : otway"
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    28
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    29
         (*Alice initiates a protocol run*)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    30
    OR1  "[| evs: otway;  A ~= B;  B ~= Server;  Nonce NA ~: used evs |]
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    31
          ==> Says A B {|Nonce NA, Agent A, Agent B, 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    32
                         Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |} 
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    33
                 # evs : otway"
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    34
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    35
         (*Bob's response to Alice's message.  Bob doesn't know who 
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    36
	   the sender is, hence the A' in the sender field.
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    37
           We modify the published protocol by NOT encrypting NB.*)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    38
    OR2  "[| evs: otway;  B ~= Server;  Nonce NB ~: used evs;
3465
e85c24717cad set_of_list -> set
nipkow
parents: 2837
diff changeset
    39
             Says A' B {|Nonce NA, Agent A, Agent B, X|} : set evs |]
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    40
          ==> Says B Server 
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    41
                  {|Nonce NA, Agent A, Agent B, X, Nonce NB,
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2131
diff changeset
    42
                    Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    43
                 # evs : otway"
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    44
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    45
         (*The Server receives Bob's message and checks that the three NAs
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    46
           match.  Then he sends a new session key to Bob with a packet for
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    47
           forwarding to Alice.*)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    48
    OR3  "[| evs: otway;  B ~= Server;  Key KAB ~: used evs;
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    49
             Says B' Server 
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    50
                  {|Nonce NA, Agent A, Agent B, 
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2131
diff changeset
    51
                    Crypt (shrK A) {|Nonce NA, Agent A, Agent B|}, 
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    52
                    Nonce NB, 
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2131
diff changeset
    53
                    Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
3465
e85c24717cad set_of_list -> set
nipkow
parents: 2837
diff changeset
    54
               : set evs |]
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    55
          ==> Says Server B 
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    56
                  {|Nonce NA, 
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    57
                    Crypt (shrK A) {|Nonce NA, Key KAB|},
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    58
                    Crypt (shrK B) {|Nonce NB, Key KAB|}|}
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    59
                 # evs : otway"
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    60
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    61
         (*Bob receives the Server's (?) message and compares the Nonces with
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    62
	   those in the message he previously sent the Server.*)
2131
3106a99d30a5 Changing from the Reveal to the Oops rule
paulson
parents: 2108
diff changeset
    63
    OR4  "[| evs: otway;  A ~= B;
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    64
             Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB, X''|}
3465
e85c24717cad set_of_list -> set
nipkow
parents: 2837
diff changeset
    65
               : set evs;
2837
dee1c7f1f576 Trivial renamings (for consistency with CSFW papers)
paulson
parents: 2516
diff changeset
    66
             Says S' B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
3465
e85c24717cad set_of_list -> set
nipkow
parents: 2837
diff changeset
    67
               : set evs |]
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    68
          ==> Says B A {|Nonce NA, X|} # evs : otway"
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    69
2131
3106a99d30a5 Changing from the Reveal to the Oops rule
paulson
parents: 2108
diff changeset
    70
         (*This message models possible leaks of session keys.  The nonces
3106a99d30a5 Changing from the Reveal to the Oops rule
paulson
parents: 2108
diff changeset
    71
           identify the protocol run.*)
3106a99d30a5 Changing from the Reveal to the Oops rule
paulson
parents: 2108
diff changeset
    72
    Oops "[| evs: otway;  B ~= Spy;
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2131
diff changeset
    73
             Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
3465
e85c24717cad set_of_list -> set
nipkow
parents: 2837
diff changeset
    74
               : set evs |]
2131
3106a99d30a5 Changing from the Reveal to the Oops rule
paulson
parents: 2108
diff changeset
    75
          ==> Says B Spy {|Nonce NA, Nonce NB, Key K|} # evs : otway"
2002
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    76
ed423882c6a9 Bad version of Otway-Rees and the new attack on it
paulson
parents:
diff changeset
    77
end