author | sandnerr |
Fri, 13 Dec 1996 12:01:26 +0100 | |
changeset 2380 | 90280b3a538b |
parent 2376 | f5c61fd9b9b6 |
child 2451 | ce85a2aafc7a |
permissions | -rw-r--r-- |
1934 | 1 |
(* Title: HOL/Auth/Shared |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
Theory of Shared Keys (common to all symmetric-key protocols) |
|
7 |
||
8 |
Server keys; initial states of agents; new nonces and keys; function "sees" |
|
9 |
*) |
|
10 |
||
11 |
Shared = Message + List + |
|
12 |
||
13 |
consts |
|
1943 | 14 |
shrK :: agent => key (*symmetric keys*) |
1967 | 15 |
|
1934 | 16 |
rules |
1943 | 17 |
isSym_shrK "isSymKey (shrK A)" |
1934 | 18 |
|
19 |
consts (*Initial states of agents -- parameter of the construction*) |
|
2032 | 20 |
initState :: [agent set, agent] => msg set |
1934 | 21 |
|
22 |
primrec initState agent |
|
2319 | 23 |
(*Server knows all long-term keys; other agents know only their own*) |
2032 | 24 |
initState_Server "initState lost Server = Key `` range shrK" |
25 |
initState_Friend "initState lost (Friend i) = {Key (shrK (Friend i))}" |
|
2064 | 26 |
initState_Spy "initState lost Spy = Key``shrK``lost" |
2032 | 27 |
|
1934 | 28 |
|
29 |
datatype (*Messages, and components of agent stores*) |
|
30 |
event = Says agent agent msg |
|
31 |
||
32 |
consts |
|
33 |
sees1 :: [agent, event] => msg set |
|
34 |
||
35 |
primrec sees1 event |
|
2078 | 36 |
(*Spy reads all traffic whether addressed to him or not*) |
37 |
sees1_Says "sees1 A (Says A' B X) = (if A:{B,Spy} then {X} else {})" |
|
1934 | 38 |
|
39 |
consts |
|
2032 | 40 |
sees :: [agent set, agent, event list] => msg set |
1934 | 41 |
|
42 |
primrec sees list |
|
2032 | 43 |
sees_Nil "sees lost A [] = initState lost A" |
44 |
sees_Cons "sees lost A (ev#evs) = sees1 A ev Un sees lost A evs" |
|
1934 | 45 |
|
46 |
||
2376
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
47 |
(*Agents generate "random" numbers for use as symmetric keys, as well as |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
48 |
nonces.*) |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
49 |
|
1934 | 50 |
consts |
2376
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
51 |
random :: "nat*nat => nat" |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
52 |
|
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
53 |
constdefs |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
54 |
newN :: event list => nat |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
55 |
"newN evs == random (length evs, 0)" |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
56 |
|
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
57 |
newK :: event list => nat |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
58 |
"newK evs == random (length evs, 1)" |
1934 | 59 |
|
60 |
rules |
|
2376
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
61 |
inj_shrK "inj shrK" |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
62 |
inj_random "inj random" |
2032 | 63 |
|
2376
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
64 |
random_neq_shrK "random ij ~= shrK A" |
f5c61fd9b9b6
Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents:
2319
diff
changeset
|
65 |
isSym_random "isSymKey (random ij)" |
1934 | 66 |
|
67 |
end |