diff -r 6e11c7bfb9c7 -r ab0a9fbed4c0 src/HOL/Auth/OtwayRees.thy --- a/src/HOL/Auth/OtwayRees.thy Mon Jul 14 12:44:09 1997 +0200 +++ b/src/HOL/Auth/OtwayRees.thy Mon Jul 14 12:47:21 1997 +0200 @@ -14,40 +14,40 @@ OtwayRees = Shared + -consts otway :: agent set => event list set -inductive "otway lost" +consts otway :: event list set +inductive "otway" intrs (*Initial trace is empty*) - Nil "[]: otway lost" + Nil "[]: otway" (*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: otway lost; B ~= Spy; - X: synth (analz (sees lost Spy evs)) |] - ==> Says Spy B X # evs : otway lost" + Fake "[| evs: otway; B ~= Spy; + X: synth (analz (sees Spy evs)) |] + ==> Says Spy B X # evs : otway" (*Alice initiates a protocol run*) - OR1 "[| evs: otway lost; A ~= B; B ~= Server; Nonce NA ~: used evs |] + OR1 "[| evs: otway; A ~= B; B ~= Server; Nonce NA ~: used evs |] ==> Says A B {|Nonce NA, Agent A, Agent B, Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |} - # evs : otway lost" + # evs : otway" (*Bob's response to Alice's message. Bob doesn't know who the sender is, hence the A' in the sender field. Note that NB is encrypted.*) - OR2 "[| evs: otway lost; B ~= Server; Nonce NB ~: used evs; + OR2 "[| evs: otway; B ~= Server; Nonce NB ~: used evs; Says A' B {|Nonce NA, Agent A, Agent B, X|} : set evs |] ==> Says B Server {|Nonce NA, Agent A, Agent B, X, Crypt (shrK B) {|Nonce NA, Nonce NB, Agent A, Agent B|}|} - # evs : otway lost" + # evs : otway" (*The Server receives Bob's message and checks that the three NAs match. Then he sends a new session key to Bob with a packet for forwarding to Alice.*) - OR3 "[| evs: otway lost; B ~= Server; Key KAB ~: used evs; + OR3 "[| evs: otway; B ~= Server; Key KAB ~: used evs; Says B' Server {|Nonce NA, Agent A, Agent B, Crypt (shrK A) {|Nonce NA, Agent A, Agent B|}, @@ -57,24 +57,24 @@ {|Nonce NA, Crypt (shrK A) {|Nonce NA, Key KAB|}, Crypt (shrK B) {|Nonce NB, Key KAB|}|} - # evs : otway lost" + # evs : otway" (*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; + OR4 "[| evs: otway; A ~= B; Says B Server {|Nonce NA, Agent A, Agent B, X', Crypt (shrK B) {|Nonce NA, Nonce NB, Agent A, Agent B|}|} : set evs; Says S' B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|} : set evs |] - ==> Says B A {|Nonce NA, X|} # evs : otway lost" + ==> Says B A {|Nonce NA, X|} # evs : otway" (*This message models possible leaks of session keys. The nonces identify the protocol run.*) - Oops "[| evs: otway lost; B ~= Spy; + Oops "[| evs: otway; B ~= Spy; Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|} : set evs |] - ==> Says B Spy {|Nonce NA, Nonce NB, Key K|} # evs : otway lost" + ==> Says B Spy {|Nonce NA, Nonce NB, Key K|} # evs : otway" end