--- a/src/HOL/Auth/NS_Shared.thy Wed Sep 11 18:00:53 1996 +0200
+++ b/src/HOL/Auth/NS_Shared.thy Wed Sep 11 18:40:55 1996 +0200
@@ -21,12 +21,12 @@
(*The enemy 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 ~= Enemy; X: synth (analz (sees Enemy evs))
- |] ==> Says Enemy B X # evs : ns_shared"
+ Fake "[| evs: ns_shared; B ~= Enemy; X: synth (analz (sees Enemy evs)) |]
+ ==> Says Enemy B X # evs : ns_shared"
(*Alice initiates a protocol run, requesting to talk to any B*)
- NS1 "[| evs: ns_shared; A ~= Server
- |] ==> Says A Server {|Agent A, Agent B, Nonce (newN evs)|} # evs
+ NS1 "[| evs: ns_shared; A ~= Server |]
+ ==> Says A Server {|Agent A, Agent B, Nonce (newN evs)|} # evs
: ns_shared"
(*Server's response to Alice's message.
@@ -34,8 +34,8 @@
Server doesn't know who the true sender is, hence the A' in
the sender field.*)
NS2 "[| evs: ns_shared; A ~= B; A ~= Server;
- Says A' Server {|Agent A, Agent B, Nonce NA|} : set_of_list evs
- |] ==> Says Server A
+ Says A' Server {|Agent A, Agent B, Nonce NA|} : set_of_list evs |]
+ ==> Says Server A
(Crypt {|Nonce NA, Agent B, Key (newK evs),
(Crypt {|Key (newK evs), Agent A|} (shrK B))|}
(shrK A)) # evs : ns_shared"
@@ -47,14 +47,14 @@
Says S A (Crypt {|Nonce NA, Agent B, Key K, X|} (shrK A))
: set_of_list evs;
A = Friend i;
- Says A Server {|Agent A, Agent B, Nonce NA|} : set_of_list evs
- |] ==> Says A B X # evs : ns_shared"
+ Says A Server {|Agent A, Agent B, Nonce NA|} : set_of_list evs |]
+ ==> Says A B X # evs : ns_shared"
(*Bob's nonce exchange. He does not know who the message came
from, but responds to A because she is mentioned inside.*)
NS4 "[| evs: ns_shared; A ~= B;
- Says A' B (Crypt {|Key K, Agent A|} (shrK B)) : set_of_list evs
- |] ==> Says B A (Crypt (Nonce (newN evs)) K) # evs : ns_shared"
+ Says A' B (Crypt {|Key K, Agent A|} (shrK B)) : set_of_list evs |]
+ ==> Says B A (Crypt (Nonce (newN evs)) K) # evs : ns_shared"
(*Alice responds with the Nonce, if she has seen the key before.
We do NOT use N-1 or similar as the Enemy cannot spoof such things.
@@ -64,7 +64,7 @@
NS5 "[| evs: ns_shared; A ~= B;
Says B' A (Crypt (Nonce N) K) : set_of_list evs;
Says S A (Crypt {|Nonce NA, Agent B, Key K, X|} (shrK A))
- : set_of_list evs
- |] ==> Says A B (Crypt {|Nonce N, Nonce N|} K) # evs : ns_shared"
+ : set_of_list evs |]
+ ==> Says A B (Crypt {|Nonce N, Nonce N|} K) # evs : ns_shared"
end
--- a/src/HOL/Auth/OtwayRees.thy Wed Sep 11 18:00:53 1996 +0200
+++ b/src/HOL/Auth/OtwayRees.thy Wed Sep 11 18:40:55 1996 +0200
@@ -21,12 +21,12 @@
(*The enemy 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: otway; B ~= Enemy; X: synth (analz (sees Enemy evs))
- |] ==> Says Enemy B X # evs : otway"
+ Fake "[| evs: otway; B ~= Enemy; X: synth (analz (sees Enemy evs)) |]
+ ==> Says Enemy B X # evs : otway"
(*Alice initiates a protocol run*)
- OR1 "[| evs: otway; A ~= B
- |] ==> Says A B {|Nonce (newN evs), Agent A, Agent B,
+ OR1 "[| evs: otway; A ~= B |]
+ ==> Says A B {|Nonce (newN evs), Agent A, Agent B,
Crypt {|Nonce (newN evs), Agent A, Agent B|}
(shrK A) |}
# evs : otway"
@@ -35,8 +35,8 @@
the sender is, hence the A' in the sender field.
We modify the published protocol by NOT encrypting NB.*)
OR2 "[| evs: otway; B ~= Server;
- Says A' B {|Nonce NA, Agent A, Agent B, X|} : set_of_list evs
- |] ==> Says B Server
+ Says A' B {|Nonce NA, Agent A, Agent B, X|} : set_of_list evs |]
+ ==> Says B Server
{|Nonce NA, Agent A, Agent B, X, Nonce (newN evs),
Crypt {|Nonce NA, Agent A, Agent B|} (shrK B)|}
# evs : otway"
@@ -50,8 +50,8 @@
Crypt {|Nonce NA, Agent A, Agent B|} (shrK A),
Nonce NB,
Crypt {|Nonce NA, Agent A, Agent B|} (shrK B)|}
- : set_of_list evs
- |] ==> Says Server B
+ : set_of_list evs |]
+ ==> Says Server B
{|Nonce NA,
Crypt {|Nonce NA, Key (newK evs)|} (shrK A),
Crypt {|Nonce NB, Key (newK evs)|} (shrK B)|}
@@ -63,15 +63,15 @@
Says S B {|Nonce NA, X, Crypt {|Nonce NB, Key K|} (shrK B)|}
: set_of_list evs;
Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB, X''|}
- : set_of_list evs
- |] ==> (Says B A {|Nonce NA, X|}) # evs : otway"
+ : set_of_list evs |]
+ ==> (Says B A {|Nonce NA, X|}) # evs : otway"
(*Alice checks her Nonce, then sends a dummy message to Bob,
using the new session key.*)
OR5 "[| evs: otway;
Says B' A {|Nonce NA, Crypt {|Nonce NA, Key K|} (shrK A)|}
: set_of_list evs;
- Says A B {|Nonce NA, Agent A, Agent B, X|} : set_of_list evs
- |] ==> Says A B (Crypt (Agent A) K) # evs : otway"
+ Says A B {|Nonce NA, Agent A, Agent B, X|} : set_of_list evs |]
+ ==> Says A B (Crypt (Agent A) K) # evs : otway"
end