13508
|
1 |
(******************************************************************************
|
|
2 |
date: april 2002
|
|
3 |
author: Frederic Blanqui
|
|
4 |
email: blanqui@lri.fr
|
|
5 |
webpage: http://www.lri.fr/~blanqui/
|
|
6 |
|
|
7 |
University of Cambridge, Computer Laboratory
|
|
8 |
William Gates Building, JJ Thomson Avenue
|
|
9 |
Cambridge CB3 0FD, United Kingdom
|
|
10 |
******************************************************************************)
|
|
11 |
|
|
12 |
header{*Other Protocol-Independent Results*}
|
|
13 |
|
|
14 |
theory Proto = Guard_Public:
|
|
15 |
|
|
16 |
subsection{*protocols*}
|
|
17 |
|
|
18 |
types rule = "event set * event"
|
|
19 |
|
|
20 |
syntax msg' :: "rule => msg"
|
|
21 |
|
|
22 |
translations "msg' R" == "msg (snd R)"
|
|
23 |
|
|
24 |
types proto = "rule set"
|
|
25 |
|
|
26 |
constdefs wdef :: "proto => bool"
|
|
27 |
"wdef p == ALL R k. R:p --> Number k:parts {msg' R}
|
|
28 |
--> Number k:parts (msg`(fst R))"
|
|
29 |
|
|
30 |
subsection{*substitutions*}
|
|
31 |
|
|
32 |
record subs =
|
|
33 |
agent :: "agent => agent"
|
|
34 |
nonce :: "nat => nat"
|
|
35 |
nb :: "nat => msg"
|
|
36 |
key :: "key => key"
|
|
37 |
|
|
38 |
consts apm :: "subs => msg => msg"
|
|
39 |
|
|
40 |
primrec
|
|
41 |
"apm s (Agent A) = Agent (agent s A)"
|
|
42 |
"apm s (Nonce n) = Nonce (nonce s n)"
|
|
43 |
"apm s (Number n) = nb s n"
|
|
44 |
"apm s (Key K) = Key (key s K)"
|
|
45 |
"apm s (Hash X) = Hash (apm s X)"
|
|
46 |
"apm s (Crypt K X) = (
|
|
47 |
if (EX A. K = pubK A) then Crypt (pubK (agent s (agt K))) (apm s X)
|
|
48 |
else if (EX A. K = priK A) then Crypt (priK (agent s (agt K))) (apm s X)
|
|
49 |
else Crypt (key s K) (apm s X))"
|
|
50 |
"apm s {|X,Y|} = {|apm s X, apm s Y|}"
|
|
51 |
|
|
52 |
lemma apm_parts: "X:parts {Y} ==> apm s X:parts {apm s Y}"
|
|
53 |
apply (erule parts.induct, simp_all, blast)
|
|
54 |
apply (erule parts.Fst)
|
|
55 |
apply (erule parts.Snd)
|
|
56 |
by (erule parts.Body)+
|
|
57 |
|
|
58 |
lemma Nonce_apm [rule_format]: "Nonce n:parts {apm s X} ==>
|
|
59 |
(ALL k. Number k:parts {X} --> Nonce n ~:parts {nb s k}) -->
|
|
60 |
(EX k. Nonce k:parts {X} & nonce s k = n)"
|
|
61 |
by (induct X, simp_all, blast)
|
|
62 |
|
|
63 |
lemma wdef_Nonce: "[| Nonce n:parts {apm s X}; R:p; msg' R = X; wdef p;
|
|
64 |
Nonce n ~:parts (apm s `(msg `(fst R))) |] ==>
|
|
65 |
(EX k. Nonce k:parts {X} & nonce s k = n)"
|
|
66 |
apply (erule Nonce_apm, unfold wdef_def)
|
|
67 |
apply (drule_tac x=R in spec, drule_tac x=k in spec, clarsimp)
|
|
68 |
apply (drule_tac x=x in bspec, simp)
|
|
69 |
apply (drule_tac Y="msg x" and s=s in apm_parts, simp)
|
|
70 |
by (blast dest: parts_parts)
|
|
71 |
|
|
72 |
consts ap :: "subs => event => event"
|
|
73 |
|
|
74 |
primrec
|
|
75 |
"ap s (Says A B X) = Says (agent s A) (agent s B) (apm s X)"
|
|
76 |
"ap s (Gets A X) = Gets (agent s A) (apm s X)"
|
|
77 |
"ap s (Notes A X) = Notes (agent s A) (apm s X)"
|
|
78 |
|
|
79 |
syntax
|
|
80 |
ap' :: "rule => msg"
|
|
81 |
apm' :: "rule => msg"
|
|
82 |
priK' :: "subs => agent => key"
|
|
83 |
pubK' :: "subs => agent => key"
|
|
84 |
|
|
85 |
translations
|
|
86 |
"ap' s R" == "ap s (snd R)"
|
|
87 |
"apm' s R" == "apm s (msg' R)"
|
|
88 |
"priK' s A" == "priK (agent s A)"
|
|
89 |
"pubK' s A" == "pubK (agent s A)"
|
|
90 |
|
|
91 |
subsection{*nonces generated by a rule*}
|
|
92 |
|
|
93 |
constdefs newn :: "rule => nat set"
|
|
94 |
"newn R == {n. Nonce n:parts {msg (snd R)} & Nonce n ~:parts (msg`(fst R))}"
|
|
95 |
|
|
96 |
lemma newn_parts: "n:newn R ==> Nonce (nonce s n):parts {apm' s R}"
|
|
97 |
by (auto simp: newn_def dest: apm_parts)
|
|
98 |
|
|
99 |
subsection{*traces generated by a protocol*}
|
|
100 |
|
|
101 |
constdefs ok :: "event list => rule => subs => bool"
|
|
102 |
"ok evs R s == ((ALL x. x:fst R --> ap s x:set evs)
|
|
103 |
& (ALL n. n:newn R --> Nonce (nonce s n) ~:used evs))"
|
|
104 |
|
|
105 |
consts tr :: "proto => event list set"
|
|
106 |
|
|
107 |
inductive "tr p" intros
|
|
108 |
|
|
109 |
Nil [intro]: "[]:tr p"
|
|
110 |
|
|
111 |
Fake [intro]: "[| evsf:tr p; X:synth (analz (spies evsf)) |]
|
|
112 |
==> Says Spy B X # evsf:tr p"
|
|
113 |
|
|
114 |
Proto [intro]: "[| evs:tr p; R:p; ok evs R s |] ==> ap' s R # evs:tr p"
|
|
115 |
|
|
116 |
subsection{*general properties*}
|
|
117 |
|
|
118 |
lemma one_step_tr [iff]: "one_step (tr p)"
|
|
119 |
apply (unfold one_step_def, clarify)
|
|
120 |
by (ind_cases "ev # evs:tr p", auto)
|
|
121 |
|
|
122 |
constdefs has_only_Says' :: "proto => bool"
|
|
123 |
"has_only_Says' p == ALL R. R:p --> is_Says (snd R)"
|
|
124 |
|
|
125 |
lemma has_only_Says'D: "[| R:p; has_only_Says' p |]
|
|
126 |
==> (EX A B X. snd R = Says A B X)"
|
|
127 |
by (unfold has_only_Says'_def is_Says_def, blast)
|
|
128 |
|
|
129 |
lemma has_only_Says_tr [simp]: "has_only_Says' p ==> has_only_Says (tr p)"
|
|
130 |
apply (unfold has_only_Says_def)
|
|
131 |
apply (rule allI, rule allI, rule impI)
|
|
132 |
apply (erule tr.induct)
|
|
133 |
apply (auto simp: has_only_Says'_def ok_def)
|
|
134 |
by (drule_tac x=a in spec, auto simp: is_Says_def)
|
|
135 |
|
|
136 |
lemma has_only_Says'_in_trD: "[| has_only_Says' p; list @ ev # evs1 \<in> tr p |]
|
|
137 |
==> (EX A B X. ev = Says A B X)"
|
|
138 |
by (drule has_only_Says_tr, auto)
|
|
139 |
|
|
140 |
lemma ok_not_used: "[| Nonce n ~:used evs; ok evs R s;
|
|
141 |
ALL x. x:fst R --> is_Says x |] ==> Nonce n ~:parts (apm s `(msg `(fst R)))"
|
|
142 |
apply (unfold ok_def, clarsimp)
|
|
143 |
apply (drule_tac x=x in spec, drule_tac x=x in spec)
|
|
144 |
by (auto simp: is_Says_def dest: Says_imp_spies not_used_not_spied parts_parts)
|
|
145 |
|
|
146 |
lemma ok_is_Says: "[| evs' @ ev # evs:tr p; ok evs R s; has_only_Says' p;
|
|
147 |
R:p; x:fst R |] ==> is_Says x"
|
|
148 |
apply (unfold ok_def is_Says_def, clarify)
|
|
149 |
apply (drule_tac x=x in spec, simp)
|
|
150 |
apply (subgoal_tac "one_step (tr p)")
|
|
151 |
apply (drule trunc, simp, drule one_step_Cons, simp)
|
|
152 |
apply (drule has_only_SaysD, simp+)
|
|
153 |
by (clarify, case_tac x, auto)
|
|
154 |
|
|
155 |
subsection{*types*}
|
|
156 |
|
|
157 |
types keyfun = "rule => subs => nat => event list => key set"
|
|
158 |
|
|
159 |
types secfun = "rule => nat => subs => key set => msg"
|
|
160 |
|
|
161 |
subsection{*introduction of a fresh guarded nonce*}
|
|
162 |
|
|
163 |
constdefs fresh :: "proto => rule => subs => nat => key set => event list
|
|
164 |
=> bool"
|
|
165 |
"fresh p R s n Ks evs == (EX evs1 evs2. evs = evs2 @ ap' s R # evs1
|
|
166 |
& Nonce n ~:used evs1 & R:p & ok evs1 R s & Nonce n:parts {apm' s R}
|
|
167 |
& apm' s R:guard n Ks)"
|
|
168 |
|
|
169 |
lemma freshD: "fresh p R s n Ks evs ==> (EX evs1 evs2.
|
|
170 |
evs = evs2 @ ap' s R # evs1 & Nonce n ~:used evs1 & R:p & ok evs1 R s
|
|
171 |
& Nonce n:parts {apm' s R} & apm' s R:guard n Ks)"
|
|
172 |
by (unfold fresh_def, blast)
|
|
173 |
|
|
174 |
lemma freshI [intro]: "[| Nonce n ~:used evs1; R:p; Nonce n:parts {apm' s R};
|
|
175 |
ok evs1 R s; apm' s R:guard n Ks |]
|
|
176 |
==> fresh p R s n Ks (list @ ap' s R # evs1)"
|
|
177 |
by (unfold fresh_def, blast)
|
|
178 |
|
|
179 |
lemma freshI': "[| Nonce n ~:used evs1; (l,r):p;
|
|
180 |
Nonce n:parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r):guard n Ks |]
|
|
181 |
==> fresh p (l,r) s n Ks (evs2 @ ap s r # evs1)"
|
|
182 |
by (drule freshI, simp+)
|
|
183 |
|
|
184 |
lemma fresh_used: "[| fresh p R' s' n Ks evs; has_only_Says' p |]
|
|
185 |
==> Nonce n:used evs"
|
|
186 |
apply (unfold fresh_def, clarify)
|
|
187 |
apply (drule has_only_Says'D)
|
|
188 |
by (auto intro: parts_used_app)
|
|
189 |
|
|
190 |
lemma fresh_newn: "[| evs' @ ap' s R # evs:tr p; wdef p; has_only_Says' p;
|
|
191 |
Nonce n ~:used evs; R:p; ok evs R s; Nonce n:parts {apm' s R} |]
|
|
192 |
==> EX k. k:newn R & nonce s k = n"
|
|
193 |
apply (drule wdef_Nonce, simp+)
|
|
194 |
apply (frule ok_not_used, simp+)
|
|
195 |
apply (clarify, erule ok_is_Says, simp+)
|
|
196 |
apply (clarify, rule_tac x=k in exI, simp add: newn_def)
|
|
197 |
apply (clarify, drule_tac Y="msg x" and s=s in apm_parts)
|
|
198 |
apply (drule ok_not_used, simp+)
|
13601
|
199 |
by (clarify, erule ok_is_Says, simp+)
|
13508
|
200 |
|
|
201 |
lemma fresh_rule: "[| evs' @ ev # evs:tr p; wdef p; Nonce n ~:used evs;
|
|
202 |
Nonce n:parts {msg ev} |] ==> EX R s. R:p & ap' s R = ev"
|
|
203 |
apply (drule trunc, simp, ind_cases "ev # evs:tr p", simp)
|
|
204 |
by (drule_tac x=X in in_sub, drule parts_sub, simp, simp, blast+)
|
|
205 |
|
|
206 |
lemma fresh_ruleD: "[| fresh p R' s' n Ks evs; keys R' s' n evs <= Ks; wdef p;
|
|
207 |
has_only_Says' p; evs:tr p; ALL R k s. nonce s k = n --> Nonce n:used evs -->
|
|
208 |
R:p --> k:newn R --> Nonce n:parts {apm' s R} --> apm' s R:guard n Ks -->
|
|
209 |
apm' s R:parts (spies evs) --> keys R s n evs <= Ks --> P |] ==> P"
|
|
210 |
apply (frule fresh_used, simp)
|
|
211 |
apply (unfold fresh_def, clarify)
|
|
212 |
apply (drule_tac x=R' in spec)
|
|
213 |
apply (drule fresh_newn, simp+, clarify)
|
|
214 |
apply (drule_tac x=k in spec)
|
|
215 |
apply (drule_tac x=s' in spec)
|
|
216 |
apply (subgoal_tac "apm' s' R':parts (spies (evs2 @ ap' s' R' # evs1))")
|
|
217 |
apply (case_tac R', drule has_only_Says'D, simp, clarsimp)
|
|
218 |
apply (case_tac R', drule has_only_Says'D, simp, clarsimp)
|
|
219 |
apply (rule_tac Y="apm s' X" in parts_parts, blast)
|
|
220 |
by (rule parts.Inj, rule Says_imp_spies, simp, blast)
|
|
221 |
|
|
222 |
subsection{*safe keys*}
|
|
223 |
|
|
224 |
constdefs safe :: "key set => msg set => bool"
|
|
225 |
"safe Ks G == ALL K. K:Ks --> Key K ~:analz G"
|
|
226 |
|
|
227 |
lemma safeD [dest]: "[| safe Ks G; K:Ks |] ==> Key K ~:analz G"
|
|
228 |
by (unfold safe_def, blast)
|
|
229 |
|
|
230 |
lemma safe_insert: "safe Ks (insert X G) ==> safe Ks G"
|
|
231 |
by (unfold safe_def, blast)
|
|
232 |
|
|
233 |
lemma Guard_safe: "[| Guard n Ks G; safe Ks G |] ==> Nonce n ~:analz G"
|
|
234 |
by (blast dest: Guard_invKey)
|
|
235 |
|
|
236 |
subsection{*guardedness preservation*}
|
|
237 |
|
|
238 |
constdefs preserv :: "proto => keyfun => nat => key set => bool"
|
|
239 |
"preserv p keys n Ks == (ALL evs R' s' R s. evs:tr p -->
|
|
240 |
Guard n Ks (spies evs) --> safe Ks (spies evs) --> fresh p R' s' n Ks evs -->
|
|
241 |
keys R' s' n evs <= Ks --> R:p --> ok evs R s --> apm' s R:guard n Ks)"
|
|
242 |
|
|
243 |
lemma preservD: "[| preserv p keys n Ks; evs:tr p; Guard n Ks (spies evs);
|
|
244 |
safe Ks (spies evs); fresh p R' s' n Ks evs; R:p; ok evs R s;
|
|
245 |
keys R' s' n evs <= Ks |] ==> apm' s R:guard n Ks"
|
|
246 |
by (unfold preserv_def, blast)
|
|
247 |
|
|
248 |
lemma preservD': "[| preserv p keys n Ks; evs:tr p; Guard n Ks (spies evs);
|
|
249 |
safe Ks (spies evs); fresh p R' s' n Ks evs; (l,Says A B X):p;
|
|
250 |
ok evs (l,Says A B X) s; keys R' s' n evs <= Ks |] ==> apm s X:guard n Ks"
|
|
251 |
by (drule preservD, simp+)
|
|
252 |
|
|
253 |
subsection{*monotonic keyfun*}
|
|
254 |
|
|
255 |
constdefs monoton :: "proto => keyfun => bool"
|
|
256 |
"monoton p keys == ALL R' s' n ev evs. ev # evs:tr p -->
|
|
257 |
keys R' s' n evs <= keys R' s' n (ev # evs)"
|
|
258 |
|
|
259 |
lemma monotonD [dest]: "[| keys R' s' n (ev # evs) <= Ks; monoton p keys;
|
|
260 |
ev # evs:tr p |] ==> keys R' s' n evs <= Ks"
|
|
261 |
by (unfold monoton_def, blast)
|
|
262 |
|
|
263 |
subsection{*guardedness theorem*}
|
|
264 |
|
|
265 |
lemma Guard_tr [rule_format]: "[| evs:tr p; has_only_Says' p;
|
|
266 |
preserv p keys n Ks; monoton p keys; Guard n Ks (initState Spy) |] ==>
|
|
267 |
safe Ks (spies evs) --> fresh p R' s' n Ks evs --> keys R' s' n evs <= Ks -->
|
|
268 |
Guard n Ks (spies evs)"
|
|
269 |
apply (erule tr.induct)
|
|
270 |
(* Nil *)
|
|
271 |
apply simp
|
|
272 |
(* Fake *)
|
|
273 |
apply (clarify, drule freshD, clarsimp)
|
|
274 |
apply (case_tac evs2)
|
|
275 |
(* evs2 = [] *)
|
|
276 |
apply (frule has_only_Says'D, simp)
|
|
277 |
apply (clarsimp, blast)
|
|
278 |
(* evs2 = aa # list *)
|
|
279 |
apply (clarsimp, rule conjI)
|
|
280 |
apply (blast dest: safe_insert)
|
|
281 |
(* X:guard n Ks *)
|
|
282 |
apply (rule in_synth_Guard, simp, rule Guard_analz)
|
|
283 |
apply (blast dest: safe_insert)
|
|
284 |
apply (drule safe_insert, simp add: safe_def)
|
|
285 |
(* Proto *)
|
|
286 |
apply (clarify, drule freshD, clarify)
|
|
287 |
apply (case_tac evs2)
|
|
288 |
(* evs2 = [] *)
|
|
289 |
apply (frule has_only_Says'D, simp)
|
|
290 |
apply (frule_tac R=R' in has_only_Says'D, simp)
|
|
291 |
apply (case_tac R', clarsimp, blast)
|
|
292 |
(* evs2 = ab # list *)
|
|
293 |
apply (frule has_only_Says'D, simp)
|
|
294 |
apply (clarsimp, rule conjI)
|
|
295 |
apply (drule Proto, simp+, blast dest: safe_insert)
|
|
296 |
(* apm s X:guard n Ks *)
|
|
297 |
apply (frule Proto, simp+)
|
|
298 |
apply (erule preservD', simp+)
|
|
299 |
apply (blast dest: safe_insert)
|
|
300 |
apply (blast dest: safe_insert)
|
|
301 |
by (blast, simp, simp, blast)
|
|
302 |
|
|
303 |
subsection{*useful properties for guardedness*}
|
|
304 |
|
|
305 |
lemma newn_neq_used: "[| Nonce n:used evs; ok evs R s; k:newn R |]
|
|
306 |
==> n ~= nonce s k"
|
|
307 |
by (auto simp: ok_def)
|
|
308 |
|
|
309 |
lemma ok_Guard: "[| ok evs R s; Guard n Ks (spies evs); x:fst R; is_Says x |]
|
|
310 |
==> apm s (msg x):parts (spies evs) & apm s (msg x):guard n Ks"
|
|
311 |
apply (unfold ok_def is_Says_def, clarify)
|
|
312 |
apply (drule_tac x="Says A B X" in spec, simp)
|
|
313 |
by (drule Says_imp_spies, auto intro: parts_parts)
|
|
314 |
|
|
315 |
lemma ok_parts_not_new: "[| Y:parts (spies evs); Nonce (nonce s n):parts {Y};
|
|
316 |
ok evs R s |] ==> n ~:newn R"
|
|
317 |
by (auto simp: ok_def dest: not_used_not_spied parts_parts)
|
|
318 |
|
|
319 |
subsection{*unicity*}
|
|
320 |
|
|
321 |
constdefs uniq :: "proto => secfun => bool"
|
|
322 |
"uniq p secret == ALL evs R R' n n' Ks s s'. R:p --> R':p -->
|
|
323 |
n:newn R --> n':newn R' --> nonce s n = nonce s' n' -->
|
|
324 |
Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} -->
|
|
325 |
apm' s R:guard (nonce s n) Ks --> apm' s' R':guard (nonce s n) Ks -->
|
|
326 |
evs:tr p --> Nonce (nonce s n) ~:analz (spies evs) -->
|
|
327 |
secret R n s Ks:parts (spies evs) --> secret R' n' s' Ks:parts (spies evs) -->
|
|
328 |
secret R n s Ks = secret R' n' s' Ks"
|
|
329 |
|
|
330 |
lemma uniqD: "[| uniq p secret; evs: tr p; R:p; R':p; n:newn R; n':newn R';
|
|
331 |
nonce s n = nonce s' n'; Nonce (nonce s n) ~:analz (spies evs);
|
|
332 |
Nonce (nonce s n):parts {apm' s R}; Nonce (nonce s n):parts {apm' s' R'};
|
|
333 |
secret R n s Ks:parts (spies evs); secret R' n' s' Ks:parts (spies evs);
|
|
334 |
apm' s R:guard (nonce s n) Ks; apm' s' R':guard (nonce s n) Ks |] ==>
|
|
335 |
secret R n s Ks = secret R' n' s' Ks"
|
|
336 |
by (unfold uniq_def, blast)
|
|
337 |
|
|
338 |
constdefs ord :: "proto => (rule => rule => bool) => bool"
|
|
339 |
"ord p inf == ALL R R'. R:p --> R':p --> ~ inf R R' --> inf R' R"
|
|
340 |
|
|
341 |
lemma ordD: "[| ord p inf; ~ inf R R'; R:p; R':p |] ==> inf R' R"
|
|
342 |
by (unfold ord_def, blast)
|
|
343 |
|
|
344 |
constdefs uniq' :: "proto => (rule => rule => bool) => secfun => bool"
|
|
345 |
"uniq' p inf secret == ALL evs R R' n n' Ks s s'. R:p --> R':p -->
|
|
346 |
inf R R' --> n:newn R --> n':newn R' --> nonce s n = nonce s' n' -->
|
|
347 |
Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} -->
|
|
348 |
apm' s R:guard (nonce s n) Ks --> apm' s' R':guard (nonce s n) Ks -->
|
|
349 |
evs:tr p --> Nonce (nonce s n) ~:analz (spies evs) -->
|
|
350 |
secret R n s Ks:parts (spies evs) --> secret R' n' s' Ks:parts (spies evs) -->
|
|
351 |
secret R n s Ks = secret R' n' s' Ks"
|
|
352 |
|
|
353 |
lemma uniq'D: "[| uniq' p inf secret; evs: tr p; inf R R'; R:p; R':p; n:newn R;
|
|
354 |
n':newn R'; nonce s n = nonce s' n'; Nonce (nonce s n) ~:analz (spies evs);
|
|
355 |
Nonce (nonce s n):parts {apm' s R}; Nonce (nonce s n):parts {apm' s' R'};
|
|
356 |
secret R n s Ks:parts (spies evs); secret R' n' s' Ks:parts (spies evs);
|
|
357 |
apm' s R:guard (nonce s n) Ks; apm' s' R':guard (nonce s n) Ks |] ==>
|
|
358 |
secret R n s Ks = secret R' n' s' Ks"
|
|
359 |
by (unfold uniq'_def, blast)
|
|
360 |
|
|
361 |
lemma uniq'_imp_uniq: "[| uniq' p inf secret; ord p inf |] ==> uniq p secret"
|
|
362 |
apply (unfold uniq_def)
|
|
363 |
apply (rule allI)+
|
|
364 |
apply (case_tac "inf R R'")
|
|
365 |
apply (blast dest: uniq'D)
|
|
366 |
by (auto dest: ordD uniq'D intro: sym)
|
|
367 |
|
|
368 |
subsection{*Needham-Schroeder-Lowe*}
|
|
369 |
|
|
370 |
constdefs
|
|
371 |
a :: agent "a == Friend 0"
|
|
372 |
b :: agent "b == Friend 1"
|
|
373 |
a' :: agent "a' == Friend 2"
|
|
374 |
b' :: agent "b' == Friend 3"
|
|
375 |
Na :: nat "Na == 0"
|
|
376 |
Nb :: nat "Nb == 1"
|
|
377 |
|
|
378 |
consts
|
|
379 |
ns :: proto
|
|
380 |
ns1 :: rule
|
|
381 |
ns2 :: rule
|
|
382 |
ns3 :: rule
|
|
383 |
|
|
384 |
translations
|
|
385 |
"ns1" == "({}, Says a b (Crypt (pubK b) {|Nonce Na, Agent a|}))"
|
|
386 |
|
|
387 |
"ns2" == "({Says a' b (Crypt (pubK b) {|Nonce Na, Agent a|})},
|
|
388 |
Says b a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|}))"
|
|
389 |
|
|
390 |
"ns3" == "({Says a b (Crypt (pubK b) {|Nonce Na, Agent a|}),
|
|
391 |
Says b' a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})},
|
|
392 |
Says a b (Crypt (pubK b) (Nonce Nb)))"
|
|
393 |
|
|
394 |
inductive ns intros
|
|
395 |
[iff]: "ns1:ns"
|
|
396 |
[iff]: "ns2:ns"
|
|
397 |
[iff]: "ns3:ns"
|
|
398 |
|
|
399 |
syntax
|
|
400 |
ns3a :: msg
|
|
401 |
ns3b :: msg
|
|
402 |
|
|
403 |
translations
|
|
404 |
"ns3a" => "Says a b (Crypt (pubK b) {|Nonce Na, Agent a|})"
|
|
405 |
"ns3b" => "Says b' a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})"
|
|
406 |
|
|
407 |
constdefs keys :: "keyfun"
|
|
408 |
"keys R' s' n evs == {priK' s' a, priK' s' b}"
|
|
409 |
|
|
410 |
lemma "monoton ns keys"
|
|
411 |
by (simp add: keys_def monoton_def)
|
|
412 |
|
|
413 |
constdefs secret :: "secfun"
|
|
414 |
"secret R n s Ks ==
|
|
415 |
(if R=ns1 then apm s (Crypt (pubK b) {|Nonce Na, Agent a|})
|
|
416 |
else if R=ns2 then apm s (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})
|
|
417 |
else Number 0)"
|
|
418 |
|
|
419 |
constdefs inf :: "rule => rule => bool"
|
|
420 |
"inf R R' == (R=ns1 | (R=ns2 & R'~=ns1) | (R=ns3 & R'=ns3))"
|
|
421 |
|
|
422 |
lemma inf_is_ord [iff]: "ord ns inf"
|
|
423 |
apply (unfold ord_def inf_def)
|
|
424 |
apply (rule allI)+
|
|
425 |
by (rule impI, erule ns.cases, simp_all)+
|
|
426 |
|
|
427 |
subsection{*general properties*}
|
|
428 |
|
|
429 |
lemma ns_has_only_Says' [iff]: "has_only_Says' ns"
|
|
430 |
apply (unfold has_only_Says'_def)
|
|
431 |
apply (rule allI, rule impI)
|
|
432 |
by (erule ns.cases, auto)
|
|
433 |
|
|
434 |
lemma newn_ns1 [iff]: "newn ns1 = {Na}"
|
|
435 |
by (simp add: newn_def)
|
|
436 |
|
|
437 |
lemma newn_ns2 [iff]: "newn ns2 = {Nb}"
|
|
438 |
by (auto simp: newn_def Na_def Nb_def)
|
|
439 |
|
|
440 |
lemma newn_ns3 [iff]: "newn ns3 = {}"
|
|
441 |
by (auto simp: newn_def)
|
|
442 |
|
|
443 |
lemma ns_wdef [iff]: "wdef ns"
|
|
444 |
by (auto simp: wdef_def elim: ns.cases)
|
|
445 |
|
|
446 |
subsection{*guardedness for NSL*}
|
|
447 |
|
|
448 |
lemma "uniq ns secret ==> preserv ns keys n Ks"
|
|
449 |
apply (unfold preserv_def)
|
|
450 |
apply (rule allI)+
|
|
451 |
apply (rule impI, rule impI, rule impI, rule impI, rule impI)
|
|
452 |
apply (erule fresh_ruleD, simp, simp, simp, simp)
|
|
453 |
apply (rule allI)+
|
|
454 |
apply (rule impI, rule impI, rule impI)
|
|
455 |
apply (erule ns.cases)
|
|
456 |
(* fresh with NS1 *)
|
|
457 |
apply (rule impI, rule impI, rule impI, rule impI, rule impI, rule impI)
|
|
458 |
apply (erule ns.cases)
|
|
459 |
(* NS1 *)
|
|
460 |
apply clarsimp
|
|
461 |
apply (frule newn_neq_used, simp, simp)
|
|
462 |
apply (rule No_Nonce, simp)
|
|
463 |
(* NS2 *)
|
|
464 |
apply clarsimp
|
|
465 |
apply (frule newn_neq_used, simp, simp)
|
|
466 |
apply (case_tac "nonce sa Na = nonce s Na")
|
|
467 |
apply (frule Guard_safe, simp)
|
|
468 |
apply (frule Crypt_guard_invKey, simp)
|
|
469 |
apply (frule ok_Guard, simp, simp, simp, clarsimp)
|
|
470 |
apply (frule_tac K="pubK' s b" in Crypt_guard_invKey, simp)
|
|
471 |
apply (frule_tac R=ns1 and R'=ns1 and Ks=Ks and s=sa and s'=s in uniqD, simp+)
|
|
472 |
apply (simp add: secret_def, simp add: secret_def, force, force)
|
|
473 |
apply (simp add: secret_def keys_def, blast)
|
|
474 |
apply (rule No_Nonce, simp)
|
|
475 |
(* NS3 *)
|
|
476 |
apply clarsimp
|
|
477 |
apply (case_tac "nonce sa Na = nonce s Nb")
|
|
478 |
apply (frule Guard_safe, simp)
|
|
479 |
apply (frule Crypt_guard_invKey, simp)
|
|
480 |
apply (frule_tac x=ns3b in ok_Guard, simp, simp, simp, clarsimp)
|
|
481 |
apply (frule_tac K="pubK' s a" in Crypt_guard_invKey, simp)
|
|
482 |
apply (frule_tac R=ns1 and R'=ns2 and Ks=Ks and s=sa and s'=s in uniqD, simp+)
|
|
483 |
apply (simp add: secret_def, simp add: secret_def, force, force)
|
|
484 |
apply (simp add: secret_def, rule No_Nonce, simp)
|
|
485 |
(* fresh with NS2 *)
|
|
486 |
apply (rule impI, rule impI, rule impI, rule impI, rule impI, rule impI)
|
|
487 |
apply (erule ns.cases)
|
|
488 |
(* NS1 *)
|
|
489 |
apply clarsimp
|
|
490 |
apply (frule newn_neq_used, simp, simp)
|
|
491 |
apply (rule No_Nonce, simp)
|
|
492 |
(* NS2 *)
|
|
493 |
apply clarsimp
|
|
494 |
apply (frule newn_neq_used, simp, simp)
|
|
495 |
apply (case_tac "nonce sa Nb = nonce s Na")
|
|
496 |
apply (frule Guard_safe, simp)
|
|
497 |
apply (frule Crypt_guard_invKey, simp)
|
|
498 |
apply (frule ok_Guard, simp, simp, simp, clarsimp)
|
|
499 |
apply (frule_tac K="pubK' s b" in Crypt_guard_invKey, simp)
|
|
500 |
apply (frule_tac R=ns2 and R'=ns1 and Ks=Ks and s=sa and s'=s in uniqD, simp+)
|
|
501 |
apply (simp add: secret_def, simp add: secret_def, force, force)
|
|
502 |
apply (simp add: secret_def, rule No_Nonce, simp)
|
|
503 |
(* NS3 *)
|
|
504 |
apply clarsimp
|
|
505 |
apply (case_tac "nonce sa Nb = nonce s Nb")
|
|
506 |
apply (frule Guard_safe, simp)
|
|
507 |
apply (frule Crypt_guard_invKey, simp)
|
|
508 |
apply (frule_tac x=ns3b in ok_Guard, simp, simp, simp, clarsimp)
|
|
509 |
apply (frule_tac K="pubK' s a" in Crypt_guard_invKey, simp)
|
|
510 |
apply (frule_tac R=ns2 and R'=ns2 and Ks=Ks and s=sa and s'=s in uniqD, simp+)
|
|
511 |
apply (simp add: secret_def, simp add: secret_def, force, force)
|
|
512 |
apply (simp add: secret_def keys_def, blast)
|
|
513 |
apply (rule No_Nonce, simp)
|
|
514 |
(* fresh with NS3 *)
|
|
515 |
by simp
|
|
516 |
|
|
517 |
subsection{*unicity for NSL*}
|
|
518 |
|
|
519 |
lemma "uniq' ns inf secret"
|
|
520 |
apply (unfold uniq'_def)
|
|
521 |
apply (rule allI)+
|
|
522 |
apply (rule impI, erule ns.cases)
|
|
523 |
(* R = ns1 *)
|
|
524 |
apply (rule impI, erule ns.cases)
|
|
525 |
(* R' = ns1 *)
|
|
526 |
apply (rule impI, rule impI, rule impI, rule impI)
|
|
527 |
apply (rule impI, rule impI, rule impI, rule impI)
|
|
528 |
apply (rule impI, erule tr.induct)
|
|
529 |
(* Nil *)
|
|
530 |
apply (simp add: secret_def)
|
|
531 |
(* Fake *)
|
|
532 |
apply (clarify, simp add: secret_def)
|
|
533 |
apply (drule notin_analz_insert)
|
|
534 |
apply (drule Crypt_insert_synth, simp, simp, simp)
|
|
535 |
apply (drule Crypt_insert_synth, simp, simp, simp, simp)
|
|
536 |
(* Proto *)
|
|
537 |
apply (erule_tac P="ok evsa Ra sa" in rev_mp)
|
|
538 |
apply (erule ns.cases)
|
|
539 |
(* ns1 *)
|
|
540 |
apply (clarify, simp add: secret_def)
|
|
541 |
apply (erule disjE, erule disjE, clarsimp)
|
|
542 |
apply (drule ok_parts_not_new, simp, simp, simp)
|
|
543 |
apply (clarify, drule ok_parts_not_new, simp, simp, simp)
|
|
544 |
(* ns2 *)
|
|
545 |
apply (simp add: secret_def)
|
|
546 |
(* ns3 *)
|
|
547 |
apply (simp add: secret_def)
|
|
548 |
(* R' = ns2 *)
|
|
549 |
apply (rule impI, rule impI, rule impI, rule impI)
|
|
550 |
apply (rule impI, rule impI, rule impI, rule impI)
|
|
551 |
apply (rule impI, erule tr.induct)
|
|
552 |
(* Nil *)
|
|
553 |
apply (simp add: secret_def)
|
|
554 |
(* Fake *)
|
|
555 |
apply (clarify, simp add: secret_def)
|
|
556 |
apply (drule notin_analz_insert)
|
|
557 |
apply (drule Crypt_insert_synth, simp, simp, simp)
|
|
558 |
apply (drule_tac n="nonce s' Nb" in Crypt_insert_synth, simp, simp, simp, simp)
|
|
559 |
(* Proto *)
|
|
560 |
apply (erule_tac P="ok evsa Ra sa" in rev_mp)
|
|
561 |
apply (erule ns.cases)
|
|
562 |
(* ns1 *)
|
|
563 |
apply (clarify, simp add: secret_def)
|
|
564 |
apply (drule_tac s=sa and n=Na in ok_parts_not_new, simp, simp, simp)
|
|
565 |
(* ns2 *)
|
|
566 |
apply (clarify, simp add: secret_def)
|
|
567 |
apply (drule_tac s=sa and n=Nb in ok_parts_not_new, simp, simp, simp)
|
|
568 |
(* ns3 *)
|
|
569 |
apply (simp add: secret_def)
|
|
570 |
(* R' = ns3 *)
|
|
571 |
apply simp
|
|
572 |
(* R = ns2 *)
|
|
573 |
apply (rule impI, erule ns.cases)
|
|
574 |
(* R' = ns1 *)
|
|
575 |
apply (simp only: inf_def, blast)
|
|
576 |
(* R' = ns2 *)
|
|
577 |
apply (rule impI, rule impI, rule impI, rule impI)
|
|
578 |
apply (rule impI, rule impI, rule impI, rule impI)
|
|
579 |
apply (rule impI, erule tr.induct)
|
|
580 |
(* Nil *)
|
|
581 |
apply (simp add: secret_def)
|
|
582 |
(* Fake *)
|
|
583 |
apply (clarify, simp add: secret_def)
|
|
584 |
apply (drule notin_analz_insert)
|
|
585 |
apply (drule_tac n="nonce s' Nb" in Crypt_insert_synth, simp, simp, simp)
|
|
586 |
apply (drule_tac n="nonce s' Nb" in Crypt_insert_synth, simp, simp, simp, simp)
|
|
587 |
(* Proto *)
|
|
588 |
apply (erule_tac P="ok evsa Ra sa" in rev_mp)
|
|
589 |
apply (erule ns.cases)
|
|
590 |
(* ns1 *)
|
|
591 |
apply (simp add: secret_def)
|
|
592 |
(* ns2 *)
|
|
593 |
apply (clarify, simp add: secret_def)
|
|
594 |
apply (erule disjE, erule disjE, clarsimp, clarsimp)
|
|
595 |
apply (drule_tac s=sa and n=Nb in ok_parts_not_new, simp, simp, simp)
|
|
596 |
apply (erule disjE, clarsimp)
|
|
597 |
apply (drule_tac s=sa and n=Nb in ok_parts_not_new, simp, simp, simp)
|
|
598 |
by (simp_all add: secret_def)
|
|
599 |
|
|
600 |
end
|