src/HOL/Auth/OtwayRees_AN.thy
changeset 2516 4d68fbe6378b
parent 2451 ce85a2aafc7a
child 2837 dee1c7f1f576
--- a/src/HOL/Auth/OtwayRees_AN.thy	Fri Jan 17 11:50:09 1997 +0100
+++ b/src/HOL/Auth/OtwayRees_AN.thy	Fri Jan 17 12:49:31 1997 +0100
@@ -7,6 +7,11 @@
 
 Simplified version with minimal encryption but explicit messages
 
+Note that the formalization does not even assume that nonces are fresh.
+This is because the protocol does not rely on uniqueness of nonces for
+security, only for freshness, and the proof script does not prove freshness
+properties.
+
 From page 11 of
   Abadi and Needham.  Prudent Engineering Practice for Cryptographic Protocols.
   IEEE Trans. SE 22 (1), 1996
@@ -29,36 +34,31 @@
 
          (*Alice initiates a protocol run*)
     OR1  "[| evs: otway lost;  A ~= B;  B ~= Server |]
-          ==> Says A B {|Agent A, Agent B, Nonce (newN(length evs))|}
-                 # evs : otway lost"
+          ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs : otway lost"
 
          (*Bob's response to Alice's message.  Bob doesn't know who 
 	   the sender is, hence the A' in the sender field.*)
     OR2  "[| evs: otway lost;  B ~= Server;
              Says A' B {|Agent A, Agent B, Nonce NA|} : set_of_list evs |]
-          ==> Says B Server {|Agent A, Agent B, Nonce NA, 
-                              Nonce (newN(length evs))|}
+          ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                  # evs : otway lost"
 
          (*The Server receives Bob's message.  Then he sends a new
            session key to Bob with a packet for forwarding to Alice.*)
-    OR3  "[| evs: otway lost;  B ~= Server;  A ~= B;
+    OR3  "[| evs: otway lost;  B ~= Server;  A ~= B;  Key KAB ~: used evs;
              Says B' Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                : set_of_list evs |]
           ==> Says Server B 
-               {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, 
-                                  Key(newK(length evs))|},
-                 Crypt (shrK B) {|Nonce NB, Agent A, Agent B, 
-                                  Key(newK(length evs))|}|}
+               {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key KAB|},
+                 Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key KAB|}|}
               # evs : otway lost"
 
          (*Bob receives the Server's (?) message and compares the Nonces with
 	   those in the message he previously sent the Server.*)
     OR4  "[| evs: otway lost;  A ~= B;
-             Says S B {|X, 
-                        Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|}
+             Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                : set_of_list evs;
-             Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
+             Says S B {|X, Crypt(shrK B){|Nonce NB, Agent A, Agent B, Key K|}|}
                : set_of_list evs |]
           ==> Says B A X # evs : otway lost"