src/HOL/Auth/NS_Shared.thy
changeset 5434 9b4bed3f394c
parent 5359 bd539b72d484
child 11104 f2024fed9f0c
--- a/src/HOL/Auth/NS_Shared.thy	Tue Sep 08 14:54:21 1998 +0200
+++ b/src/HOL/Auth/NS_Shared.thy	Tue Sep 08 15:17:11 1998 +0200
@@ -21,12 +21,11 @@
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evs: ns_shared;  B ~= Spy;  
-             X: synth (analz (spies evs)) |]
+    Fake "[| evs: ns_shared;  X: synth (analz (spies evs)) |]
           ==> Says Spy B X # evs : ns_shared"
 
          (*Alice initiates a protocol run, requesting to talk to any B*)
-    NS1  "[| evs1: ns_shared;  A ~= Server;  Nonce NA ~: used evs1 |]
+    NS1  "[| evs1: ns_shared;  Nonce NA ~: used evs1 |]
           ==> Says A Server {|Agent A, Agent B, Nonce NA|} # evs1
                 :  ns_shared"
 
@@ -34,7 +33,7 @@
            !! It may respond more than once to A's request !!
 	   Server doesn't know who the true sender is, hence the A' in
                the sender field.*)
-    NS2  "[| evs2: ns_shared;  A ~= B;  A ~= Server;  Key KAB ~: used evs2;
+    NS2  "[| evs2: ns_shared;  Key KAB ~: used evs2;
              Says A' Server {|Agent A, Agent B, Nonce NA|} : set evs2 |]
           ==> Says Server A 
                 (Crypt (shrK A)
@@ -43,8 +42,8 @@
                 # evs2 : ns_shared"
 
           (*We can't assume S=Server.  Agent A "remembers" her nonce.
-            Can inductively show A ~= Server*)
-    NS3  "[| evs3: ns_shared;  A ~= B;
+            Need A ~= Server because we allow messages to self.*)
+    NS3  "[| evs3: ns_shared;  A ~= Server;
              Says S A (Crypt (shrK A) {|Nonce NA, Agent B, Key K, X|}) 
                : set evs3;
              Says A Server {|Agent A, Agent B, Nonce NA|} : set evs3 |]
@@ -52,7 +51,7 @@
 
          (*Bob's nonce exchange.  He does not know who the message came
            from, but responds to A because she is mentioned inside.*)
-    NS4  "[| evs4: ns_shared;  A ~= B;  Nonce NB ~: used evs4;
+    NS4  "[| evs4: ns_shared;  Nonce NB ~: used evs4;
              Says A' B (Crypt (shrK B) {|Key K, Agent A|}) : set evs4 |]
           ==> Says B A (Crypt K (Nonce NB)) # evs4
                 : ns_shared"
@@ -62,7 +61,7 @@
            We do NOT send NB-1 or similar as the Spy cannot spoof such things.
            Letting the Spy add or subtract 1 lets him send ALL nonces.
            Instead we distinguish the messages by sending the nonce twice.*)
-    NS5  "[| evs5: ns_shared;  A ~= B;  
+    NS5  "[| evs5: ns_shared;  
              Says B' A (Crypt K (Nonce NB)) : set evs5;
              Says S  A (Crypt (shrK A) {|Nonce NA, Agent B, Key K, X|})
                : set evs5 |]