| author | mueller | 
| Mon, 12 Jan 1998 17:51:05 +0100 | |
| changeset 4562 | 7aa75c767182 | 
| parent 3683 | aafe719dff14 | 
| child 5434 | 9b4bed3f394c | 
| permissions | -rw-r--r-- | 
| 2318 | 1 | (* Title: HOL/Auth/NS_Public | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | ||
| 6 | Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol. | |
| 2538 | 7 | Version incorporating Lowe's fix (inclusion of B's identity in round 2). | 
| 2318 | 8 | *) | 
| 9 | ||
| 10 | NS_Public = Public + | |
| 11 | ||
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3466diff
changeset | 12 | consts ns_public :: event list set | 
| 2549 | 13 | |
| 2318 | 14 | inductive ns_public | 
| 15 | intrs | |
| 16 | (*Initial trace is empty*) | |
| 17 | Nil "[]: ns_public" | |
| 18 | ||
| 19 | (*The spy MAY say anything he CAN say. We do not expect him to | |
| 20 | invent new nonces here, but he can also use NS1. Common to | |
| 21 | all similar protocols.*) | |
| 22 | Fake "[| evs: ns_public; B ~= Spy; | |
| 3683 | 23 | X: synth (analz (spies evs)) |] | 
| 2318 | 24 | ==> Says Spy B X # evs : ns_public" | 
| 25 | ||
| 26 | (*Alice initiates a protocol run, sending a nonce to Bob*) | |
| 3659 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 27 | NS1 "[| evs1: ns_public; A ~= B; Nonce NA ~: used evs1 |] | 
| 2497 | 28 |           ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|})
 | 
| 3659 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 29 | # evs1 : ns_public" | 
| 2318 | 30 | |
| 31 | (*Bob responds to Alice's message with a further nonce*) | |
| 3659 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 32 | NS2 "[| evs2: ns_public; A ~= B; Nonce NB ~: used evs2; | 
| 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 33 |              Says A' B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs2 |]
 | 
| 2497 | 34 |           ==> Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})
 | 
| 3659 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 35 | # evs2 : ns_public" | 
| 2318 | 36 | |
| 37 | (*Alice proves her existence by sending NB back to Bob.*) | |
| 3659 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 38 | NS3 "[| evs3: ns_public; | 
| 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 39 |              Says A  B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs3;
 | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2497diff
changeset | 40 |              Says B' A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})
 | 
| 3659 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 41 | : set evs3 |] | 
| 
eddedfe2f3f8
Renamed "evs" to "evs1", "evs2", etc. in protocol inductive definition
 paulson parents: 
3541diff
changeset | 42 | ==> Says A B (Crypt (pubK B) (Nonce NB)) # evs3 : ns_public" | 
| 2318 | 43 | |
| 44 | (**Oops message??**) | |
| 45 | ||
| 46 | end |