author | blanchet |
Wed, 15 Dec 2010 18:10:32 +0100 | |
changeset 41171 | 043f8dc3b51f |
parent 39216 | 62332b382dba |
child 41775 | 6214816d79d3 |
permissions | -rw-r--r-- |
13508 | 1 |
(****************************************************************************** |
2 |
from G. Karjoth, N. Asokan and C. Gulcu |
|
3 |
"Protecting the computation results of free-roaming agents" |
|
4 |
Mobiles Agents 1998, LNCS 1477 |
|
5 |
||
6 |
date: february 2002 |
|
7 |
author: Frederic Blanqui |
|
8 |
email: blanqui@lri.fr |
|
9 |
webpage: http://www.lri.fr/~blanqui/ |
|
10 |
||
11 |
University of Cambridge, Computer Laboratory |
|
12 |
William Gates Building, JJ Thomson Avenue |
|
13 |
Cambridge CB3 0FD, United Kingdom |
|
14 |
******************************************************************************) |
|
15 |
||
16 |
header{*Protocol P1*} |
|
17 |
||
39216 | 18 |
theory P1 imports "../Public" Guard_Public List_Msg begin |
13508 | 19 |
|
20 |
subsection{*Protocol Definition*} |
|
21 |
||
22 |
(****************************************************************************** |
|
23 |
||
24 |
the contents of the messages are not completely specified in the paper |
|
25 |
we assume that the user sends his request and his itinerary in the clear |
|
26 |
||
27 |
we will adopt the following format for messages: {|A,r,I,L|} |
|
28 |
A: originator (agent) |
|
29 |
r: request (number) |
|
30 |
I: next shops (agent list) |
|
31 |
L: collected offers (offer list) |
|
32 |
||
33 |
in the paper, the authors use nonces r_i to add redundancy in the offer |
|
34 |
in order to make it safer against dictionary attacks |
|
35 |
it is not necessary in our modelization since crypto is assumed to be strong |
|
36 |
(Crypt in injective) |
|
37 |
******************************************************************************) |
|
38 |
||
39 |
subsubsection{*offer chaining: |
|
40 |
B chains his offer for A with the head offer of L for sending it to C*} |
|
41 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
42 |
definition chain :: "agent => nat => agent => msg => agent => msg" where |
13508 | 43 |
"chain B ofr A L C == |
44 |
let m1= Crypt (pubK A) (Nonce ofr) in |
|
45 |
let m2= Hash {|head L, Agent C|} in |
|
46 |
sign B {|m1,m2|}" |
|
47 |
||
48 |
declare Let_def [simp] |
|
49 |
||
50 |
lemma chain_inj [iff]: "(chain B ofr A L C = chain B' ofr' A' L' C') |
|
51 |
= (B=B' & ofr=ofr' & A=A' & head L = head L' & C=C')" |
|
52 |
by (auto simp: chain_def Let_def) |
|
53 |
||
54 |
lemma Nonce_in_chain [iff]: "Nonce ofr:parts {chain B ofr A L C}" |
|
55 |
by (auto simp: chain_def sign_def) |
|
56 |
||
57 |
subsubsection{*agent whose key is used to sign an offer*} |
|
58 |
||
35418 | 59 |
fun shop :: "msg => msg" where |
13508 | 60 |
"shop {|B,X,Crypt K H|} = Agent (agt K)" |
61 |
||
62 |
lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B" |
|
63 |
by (simp add: chain_def sign_def) |
|
64 |
||
65 |
subsubsection{*nonce used in an offer*} |
|
66 |
||
35418 | 67 |
fun nonce :: "msg => msg" where |
13508 | 68 |
"nonce {|B,{|Crypt K ofr,m2|},CryptH|} = ofr" |
69 |
||
70 |
lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr" |
|
71 |
by (simp add: chain_def sign_def) |
|
72 |
||
73 |
subsubsection{*next shop*} |
|
74 |
||
35418 | 75 |
fun next_shop :: "msg => agent" where |
13508 | 76 |
"next_shop {|B,{|m1,Hash{|headL,Agent C|}|},CryptH|} = C" |
77 |
||
78 |
lemma next_shop_chain [iff]: "next_shop (chain B ofr A L C) = C" |
|
79 |
by (simp add: chain_def sign_def) |
|
80 |
||
81 |
subsubsection{*anchor of the offer list*} |
|
82 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
83 |
definition anchor :: "agent => nat => agent => msg" where |
13508 | 84 |
"anchor A n B == chain A n A (cons nil nil) B" |
85 |
||
86 |
lemma anchor_inj [iff]: "(anchor A n B = anchor A' n' B') |
|
87 |
= (A=A' & n=n' & B=B')" |
|
88 |
by (auto simp: anchor_def) |
|
89 |
||
90 |
lemma Nonce_in_anchor [iff]: "Nonce n:parts {anchor A n B}" |
|
91 |
by (auto simp: anchor_def) |
|
92 |
||
93 |
lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A" |
|
94 |
by (simp add: anchor_def) |
|
95 |
||
96 |
lemma nonce_anchor [simp]: "nonce (anchor A n B) = Nonce n" |
|
97 |
by (simp add: anchor_def) |
|
98 |
||
99 |
lemma next_shop_anchor [iff]: "next_shop (anchor A n B) = B" |
|
100 |
by (simp add: anchor_def) |
|
101 |
||
102 |
subsubsection{*request event*} |
|
103 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
104 |
definition reqm :: "agent => nat => nat => msg => agent => msg" where |
13508 | 105 |
"reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I), |
106 |
cons (anchor A n B) nil|}" |
|
107 |
||
108 |
lemma reqm_inj [iff]: "(reqm A r n I B = reqm A' r' n' I' B') |
|
109 |
= (A=A' & r=r' & n=n' & I=I' & B=B')" |
|
110 |
by (auto simp: reqm_def) |
|
111 |
||
112 |
lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}" |
|
113 |
by (auto simp: reqm_def) |
|
114 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
115 |
definition req :: "agent => nat => nat => msg => agent => event" where |
13508 | 116 |
"req A r n I B == Says A B (reqm A r n I B)" |
117 |
||
118 |
lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B') |
|
119 |
= (A=A' & r=r' & n=n' & I=I' & B=B')" |
|
120 |
by (auto simp: req_def) |
|
121 |
||
122 |
subsubsection{*propose event*} |
|
123 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
124 |
definition prom :: "agent => nat => agent => nat => msg => msg => |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
125 |
msg => agent => msg" where |
13508 | 126 |
"prom B ofr A r I L J C == {|Agent A, Number r, |
127 |
app (J, del (Agent B, I)), cons (chain B ofr A L C) L|}" |
|
128 |
||
129 |
lemma prom_inj [dest]: "prom B ofr A r I L J C |
|
130 |
= prom B' ofr' A' r' I' L' J' C' |
|
131 |
==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" |
|
132 |
by (auto simp: prom_def) |
|
133 |
||
134 |
lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}" |
|
135 |
by (auto simp: prom_def) |
|
136 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
137 |
definition pro :: "agent => nat => agent => nat => msg => msg => |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
138 |
msg => agent => event" where |
13508 | 139 |
"pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)" |
140 |
||
141 |
lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C' |
|
142 |
==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" |
|
143 |
by (auto simp: pro_def dest: prom_inj) |
|
144 |
||
145 |
subsubsection{*protocol*} |
|
146 |
||
23746 | 147 |
inductive_set p1 :: "event list set" |
148 |
where |
|
13508 | 149 |
|
23746 | 150 |
Nil: "[]:p1" |
13508 | 151 |
|
23746 | 152 |
| Fake: "[| evsf:p1; X:synth (analz (spies evsf)) |] ==> Says Spy B X # evsf : p1" |
13508 | 153 |
|
23746 | 154 |
| Request: "[| evsr:p1; Nonce n ~:used evsr; I:agl |] ==> req A r n I B # evsr : p1" |
13508 | 155 |
|
23746 | 156 |
| Propose: "[| evsp:p1; Says A' B {|Agent A,Number r,I,cons M L|}:set evsp; |
157 |
I:agl; J:agl; isin (Agent C, app (J, del (Agent B, I))); |
|
158 |
Nonce ofr ~:used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp : p1" |
|
13508 | 159 |
|
160 |
subsubsection{*Composition of Traces*} |
|
161 |
||
162 |
lemma "evs':p1 ==> |
|
163 |
evs:p1 & (ALL n. Nonce n:used evs' --> Nonce n ~:used evs) --> |
|
164 |
evs'@evs : p1" |
|
165 |
apply (erule p1.induct, safe) |
|
166 |
apply (simp_all add: used_ConsI) |
|
167 |
apply (erule p1.Fake, erule synth_sub, rule analz_mono, rule knows_sub_app) |
|
168 |
apply (erule p1.Request, safe, simp_all add: req_def, force) |
|
169 |
apply (erule_tac A'=A' in p1.Propose, simp_all) |
|
170 |
apply (drule_tac x=ofr in spec, simp add: pro_def, blast) |
|
171 |
apply (erule_tac A'=A' in p1.Propose, auto simp: pro_def) |
|
172 |
done |
|
173 |
||
174 |
subsubsection{*Valid Offer Lists*} |
|
175 |
||
23746 | 176 |
inductive_set |
177 |
valid :: "agent => nat => agent => msg set" |
|
178 |
for A :: agent and n :: nat and B :: agent |
|
179 |
where |
|
180 |
Request [intro]: "cons (anchor A n B) nil:valid A n B" |
|
13508 | 181 |
|
23746 | 182 |
| Propose [intro]: "L:valid A n B |
13508 | 183 |
==> cons (chain (next_shop (head L)) ofr A L C) L:valid A n B" |
184 |
||
185 |
subsubsection{*basic properties of valid*} |
|
186 |
||
187 |
lemma valid_not_empty: "L:valid A n B ==> EX M L'. L = cons M L'" |
|
188 |
by (erule valid.cases, auto) |
|
189 |
||
190 |
lemma valid_pos_len: "L:valid A n B ==> 0 < len L" |
|
191 |
by (erule valid.induct, auto) |
|
192 |
||
193 |
subsubsection{*offers of an offer list*} |
|
194 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
195 |
definition offer_nonces :: "msg => msg set" where |
13508 | 196 |
"offer_nonces L == {X. X:parts {L} & (EX n. X = Nonce n)}" |
197 |
||
198 |
subsubsection{*the originator can get the offers*} |
|
199 |
||
200 |
lemma "L:valid A n B ==> offer_nonces L <= analz (insert L (initState A))" |
|
201 |
by (erule valid.induct, auto simp: anchor_def chain_def sign_def |
|
202 |
offer_nonces_def initState.simps) |
|
203 |
||
204 |
subsubsection{*list of offers*} |
|
205 |
||
35418 | 206 |
fun offers :: "msg => msg" where |
207 |
"offers (cons M L) = cons {|shop M, nonce M|} (offers L)" | |
|
13508 | 208 |
"offers other = nil" |
209 |
||
210 |
subsubsection{*list of agents whose keys are used to sign a list of offers*} |
|
211 |
||
35418 | 212 |
fun shops :: "msg => msg" where |
213 |
"shops (cons M L) = cons (shop M) (shops L)" | |
|
13508 | 214 |
"shops other = other" |
215 |
||
216 |
lemma shops_in_agl: "L:valid A n B ==> shops L:agl" |
|
217 |
by (erule valid.induct, auto simp: anchor_def chain_def sign_def) |
|
218 |
||
219 |
subsubsection{*builds a trace from an itinerary*} |
|
220 |
||
35418 | 221 |
fun offer_list :: "agent * nat * agent * msg * nat => msg" where |
222 |
"offer_list (A,n,B,nil,ofr) = cons (anchor A n B) nil" | |
|
13508 | 223 |
"offer_list (A,n,B,cons (Agent C) I,ofr) = ( |
224 |
let L = offer_list (A,n,B,I,Suc ofr) in |
|
225 |
cons (chain (next_shop (head L)) ofr A L C) L)" |
|
226 |
||
227 |
lemma "I:agl ==> ALL ofr. offer_list (A,n,B,I,ofr):valid A n B" |
|
228 |
by (erule agl.induct, auto) |
|
229 |
||
35418 | 230 |
fun trace :: "agent * nat * agent * nat * msg * msg * msg |
231 |
=> event list" where |
|
232 |
"trace (B,ofr,A,r,I,L,nil) = []" | |
|
13508 | 233 |
"trace (B,ofr,A,r,I,L,cons (Agent D) K) = ( |
234 |
let C = (if K=nil then B else agt_nb (head K)) in |
|
235 |
let I' = (if K=nil then cons (Agent A) (cons (Agent B) I) |
|
236 |
else cons (Agent A) (app (I, cons (head K) nil))) in |
|
237 |
let I'' = app (I, cons (head K) nil) in |
|
238 |
pro C (Suc ofr) A r I' L nil D |
|
239 |
# trace (B,Suc ofr,A,r,I'',tail L,K))" |
|
240 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
241 |
definition trace' :: "agent => nat => nat => msg => agent => nat => event list" where |
13508 | 242 |
"trace' A r n I B ofr == ( |
243 |
let AI = cons (Agent A) I in |
|
244 |
let L = offer_list (A,n,B,AI,ofr) in |
|
245 |
trace (B,ofr,A,r,nil,L,AI))" |
|
246 |
||
247 |
declare trace'_def [simp] |
|
248 |
||
249 |
subsubsection{*there is a trace in which the originator receives a valid answer*} |
|
250 |
||
251 |
lemma p1_not_empty: "evs:p1 ==> req A r n I B:set evs --> |
|
252 |
(EX evs'. evs'@evs:p1 & pro B' ofr A r I' L J A:set evs' & L:valid A n B)" |
|
253 |
oops |
|
254 |
||
255 |
||
256 |
subsection{*properties of protocol P1*} |
|
257 |
||
258 |
text{*publicly verifiable forward integrity: |
|
259 |
anyone can verify the validity of an offer list*} |
|
260 |
||
261 |
subsubsection{*strong forward integrity: |
|
262 |
except the last one, no offer can be modified*} |
|
263 |
||
264 |
lemma strong_forward_integrity: "ALL L. Suc i < len L |
|
265 |
--> L:valid A n B & repl (L,Suc i,M):valid A n B --> M = ith (L,Suc i)" |
|
266 |
apply (induct i) |
|
267 |
(* i = 0 *) |
|
268 |
apply clarify |
|
269 |
apply (frule len_not_empty, clarsimp) |
|
270 |
apply (frule len_not_empty, clarsimp) |
|
23746 | 271 |
apply (ind_cases "{|x,xa,l'a|}:valid A n B" for x xa l'a) |
272 |
apply (ind_cases "{|x,M,l'a|}:valid A n B" for x l'a) |
|
13508 | 273 |
apply (simp add: chain_def) |
274 |
(* i > 0 *) |
|
275 |
apply clarify |
|
276 |
apply (frule len_not_empty, clarsimp) |
|
23746 | 277 |
apply (ind_cases "{|x,repl(l',Suc na,M)|}:valid A n B" for x l' na) |
13508 | 278 |
apply (frule len_not_empty, clarsimp) |
23746 | 279 |
apply (ind_cases "{|x,l'|}:valid A n B" for x l') |
13508 | 280 |
by (drule_tac x=l' in spec, simp, blast) |
281 |
||
282 |
subsubsection{*insertion resilience: |
|
283 |
except at the beginning, no offer can be inserted*} |
|
284 |
||
285 |
lemma chain_isnt_head [simp]: "L:valid A n B ==> |
|
286 |
head L ~= chain (next_shop (head L)) ofr A L C" |
|
287 |
by (erule valid.induct, auto simp: chain_def sign_def anchor_def) |
|
288 |
||
289 |
lemma insertion_resilience: "ALL L. L:valid A n B --> Suc i < len L |
|
290 |
--> ins (L,Suc i,M) ~:valid A n B" |
|
291 |
apply (induct i) |
|
292 |
(* i = 0 *) |
|
293 |
apply clarify |
|
294 |
apply (frule len_not_empty, clarsimp) |
|
23746 | 295 |
apply (ind_cases "{|x,l'|}:valid A n B" for x l', simp) |
296 |
apply (ind_cases "{|x,M,l'|}:valid A n B" for x l', clarsimp) |
|
297 |
apply (ind_cases "{|head l',l'|}:valid A n B" for l', simp, simp) |
|
13508 | 298 |
(* i > 0 *) |
299 |
apply clarify |
|
300 |
apply (frule len_not_empty, clarsimp) |
|
23746 | 301 |
apply (ind_cases "{|x,l'|}:valid A n B" for x l') |
13508 | 302 |
apply (frule len_not_empty, clarsimp) |
23746 | 303 |
apply (ind_cases "{|x,ins(l',Suc na,M)|}:valid A n B" for x l' na) |
13508 | 304 |
apply (frule len_not_empty, clarsimp) |
305 |
by (drule_tac x=l' in spec, clarsimp) |
|
306 |
||
307 |
subsubsection{*truncation resilience: |
|
308 |
only shop i can truncate at offer i*} |
|
309 |
||
310 |
lemma truncation_resilience: "ALL L. L:valid A n B --> Suc i < len L |
|
311 |
--> cons M (trunc (L,Suc i)):valid A n B --> shop M = shop (ith (L,i))" |
|
312 |
apply (induct i) |
|
313 |
(* i = 0 *) |
|
314 |
apply clarify |
|
315 |
apply (frule len_not_empty, clarsimp) |
|
23746 | 316 |
apply (ind_cases "{|x,l'|}:valid A n B" for x l') |
13508 | 317 |
apply (frule len_not_empty, clarsimp) |
23746 | 318 |
apply (ind_cases "{|M,l'|}:valid A n B" for l') |
13508 | 319 |
apply (frule len_not_empty, clarsimp, simp) |
320 |
(* i > 0 *) |
|
321 |
apply clarify |
|
322 |
apply (frule len_not_empty, clarsimp) |
|
23746 | 323 |
apply (ind_cases "{|x,l'|}:valid A n B" for x l') |
13508 | 324 |
apply (frule len_not_empty, clarsimp) |
325 |
by (drule_tac x=l' in spec, clarsimp) |
|
326 |
||
327 |
subsubsection{*declarations for tactics*} |
|
328 |
||
329 |
declare knows_Spy_partsEs [elim] |
|
330 |
declare Fake_parts_insert [THEN subsetD, dest] |
|
331 |
declare initState.simps [simp del] |
|
332 |
||
333 |
subsubsection{*get components of a message*} |
|
334 |
||
335 |
lemma get_ML [dest]: "Says A' B {|A,r,I,M,L|}:set evs ==> |
|
336 |
M:parts (spies evs) & L:parts (spies evs)" |
|
337 |
by blast |
|
338 |
||
339 |
subsubsection{*general properties of p1*} |
|
340 |
||
341 |
lemma reqm_neq_prom [iff]: |
|
342 |
"reqm A r n I B ~= prom B' ofr A' r' I' (cons M L) J C" |
|
343 |
by (auto simp: reqm_def prom_def) |
|
344 |
||
345 |
lemma prom_neq_reqm [iff]: |
|
346 |
"prom B' ofr A' r' I' (cons M L) J C ~= reqm A r n I B" |
|
347 |
by (auto simp: reqm_def prom_def) |
|
348 |
||
349 |
lemma req_neq_pro [iff]: "req A r n I B ~= pro B' ofr A' r' I' (cons M L) J C" |
|
350 |
by (auto simp: req_def pro_def) |
|
351 |
||
352 |
lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C ~= req A r n I B" |
|
353 |
by (auto simp: req_def pro_def) |
|
354 |
||
355 |
lemma p1_has_no_Gets: "evs:p1 ==> ALL A X. Gets A X ~:set evs" |
|
356 |
by (erule p1.induct, auto simp: req_def pro_def) |
|
357 |
||
358 |
lemma p1_is_Gets_correct [iff]: "Gets_correct p1" |
|
359 |
by (auto simp: Gets_correct_def dest: p1_has_no_Gets) |
|
360 |
||
361 |
lemma p1_is_one_step [iff]: "one_step p1" |
|
23746 | 362 |
by (unfold one_step_def, clarify, ind_cases "ev#evs:p1" for ev evs, auto) |
13508 | 363 |
|
364 |
lemma p1_has_only_Says' [rule_format]: "evs:p1 ==> |
|
365 |
ev:set evs --> (EX A B X. ev=Says A B X)" |
|
366 |
by (erule p1.induct, auto simp: req_def pro_def) |
|
367 |
||
368 |
lemma p1_has_only_Says [iff]: "has_only_Says p1" |
|
369 |
by (auto simp: has_only_Says_def dest: p1_has_only_Says') |
|
370 |
||
371 |
lemma p1_is_regular [iff]: "regular p1" |
|
372 |
apply (simp only: regular_def, clarify) |
|
373 |
apply (erule_tac p1.induct) |
|
374 |
apply (simp_all add: initState.simps knows.simps pro_def prom_def |
|
375 |
req_def reqm_def anchor_def chain_def sign_def) |
|
376 |
by (auto dest: no_Key_in_agl no_Key_in_appdel parts_trans) |
|
377 |
||
378 |
subsubsection{*private keys are safe*} |
|
379 |
||
380 |
lemma priK_parts_Friend_imp_bad [rule_format,dest]: |
|
381 |
"[| evs:p1; Friend B ~= A |] |
|
382 |
==> (Key (priK A):parts (knows (Friend B) evs)) --> (A:bad)" |
|
383 |
apply (erule p1.induct) |
|
384 |
apply (simp_all add: initState.simps knows.simps pro_def prom_def |
|
17778 | 385 |
req_def reqm_def anchor_def chain_def sign_def) |
13508 | 386 |
apply (blast dest: no_Key_in_agl) |
387 |
apply (auto del: parts_invKey disjE dest: parts_trans |
|
388 |
simp add: no_Key_in_appdel) |
|
389 |
done |
|
390 |
||
391 |
lemma priK_analz_Friend_imp_bad [rule_format,dest]: |
|
392 |
"[| evs:p1; Friend B ~= A |] |
|
393 |
==> (Key (priK A):analz (knows (Friend B) evs)) --> (A:bad)" |
|
394 |
by auto |
|
395 |
||
396 |
lemma priK_notin_knows_max_Friend: "[| evs:p1; A ~:bad; A ~= Friend C |] |
|
397 |
==> Key (priK A) ~:analz (knows_max (Friend C) evs)" |
|
398 |
apply (rule not_parts_not_analz, simp add: knows_max_def, safe) |
|
399 |
apply (drule_tac H="spies' evs" in parts_sub) |
|
400 |
apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) |
|
401 |
apply (drule_tac H="spies evs" in parts_sub) |
|
402 |
by (auto dest: knows'_sub_knows [THEN subsetD] priK_notin_initState_Friend) |
|
403 |
||
404 |
subsubsection{*general guardedness properties*} |
|
405 |
||
406 |
lemma agl_guard [intro]: "I:agl ==> I:guard n Ks" |
|
407 |
by (erule agl.induct, auto) |
|
408 |
||
409 |
lemma Says_to_knows_max'_guard: "[| Says A' C {|A'',r,I,L|}:set evs; |
|
410 |
Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks" |
|
411 |
by (auto dest: Says_to_knows_max') |
|
412 |
||
413 |
lemma Says_from_knows_max'_guard: "[| Says C A' {|A'',r,I,L|}:set evs; |
|
414 |
Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks" |
|
415 |
by (auto dest: Says_from_knows_max') |
|
416 |
||
417 |
lemma Says_Nonce_not_used_guard: "[| Says A' B {|A'',r,I,L|}:set evs; |
|
418 |
Nonce n ~:used evs |] ==> L:guard n Ks" |
|
419 |
by (drule not_used_not_parts, auto) |
|
420 |
||
421 |
subsubsection{*guardedness of messages*} |
|
422 |
||
423 |
lemma chain_guard [iff]: "chain B ofr A L C:guard n {priK A}" |
|
424 |
by (case_tac "ofr=n", auto simp: chain_def sign_def) |
|
425 |
||
426 |
lemma chain_guard_Nonce_neq [intro]: "n ~= ofr |
|
427 |
==> chain B ofr A' L C:guard n {priK A}" |
|
428 |
by (auto simp: chain_def sign_def) |
|
429 |
||
430 |
lemma anchor_guard [iff]: "anchor A n' B:guard n {priK A}" |
|
431 |
by (case_tac "n'=n", auto simp: anchor_def) |
|
432 |
||
433 |
lemma anchor_guard_Nonce_neq [intro]: "n ~= n' |
|
434 |
==> anchor A' n' B:guard n {priK A}" |
|
435 |
by (auto simp: anchor_def) |
|
436 |
||
437 |
lemma reqm_guard [intro]: "I:agl ==> reqm A r n' I B:guard n {priK A}" |
|
438 |
by (case_tac "n'=n", auto simp: reqm_def) |
|
439 |
||
440 |
lemma reqm_guard_Nonce_neq [intro]: "[| n ~= n'; I:agl |] |
|
441 |
==> reqm A' r n' I B:guard n {priK A}" |
|
442 |
by (auto simp: reqm_def) |
|
443 |
||
444 |
lemma prom_guard [intro]: "[| I:agl; J:agl; L:guard n {priK A} |] |
|
445 |
==> prom B ofr A r I L J C:guard n {priK A}" |
|
446 |
by (auto simp: prom_def) |
|
447 |
||
448 |
lemma prom_guard_Nonce_neq [intro]: "[| n ~= ofr; I:agl; J:agl; |
|
449 |
L:guard n {priK A} |] ==> prom B ofr A' r I L J C:guard n {priK A}" |
|
450 |
by (auto simp: prom_def) |
|
451 |
||
452 |
subsubsection{*Nonce uniqueness*} |
|
453 |
||
454 |
lemma uniq_Nonce_in_chain [dest]: "Nonce k:parts {chain B ofr A L C} ==> k=ofr" |
|
455 |
by (auto simp: chain_def sign_def) |
|
456 |
||
457 |
lemma uniq_Nonce_in_anchor [dest]: "Nonce k:parts {anchor A n B} ==> k=n" |
|
458 |
by (auto simp: anchor_def chain_def sign_def) |
|
459 |
||
460 |
lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k:parts {reqm A r n I B}; |
|
461 |
I:agl |] ==> k=n" |
|
462 |
by (auto simp: reqm_def dest: no_Nonce_in_agl) |
|
463 |
||
464 |
lemma uniq_Nonce_in_prom [dest]: "[| Nonce k:parts {prom B ofr A r I L J C}; |
|
465 |
I:agl; J:agl; Nonce k ~:parts {L} |] ==> k=ofr" |
|
466 |
by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel) |
|
467 |
||
468 |
subsubsection{*requests are guarded*} |
|
469 |
||
470 |
lemma req_imp_Guard [rule_format]: "[| evs:p1; A ~:bad |] ==> |
|
471 |
req A r n I B:set evs --> Guard n {priK A} (spies evs)" |
|
472 |
apply (erule p1.induct, simp) |
|
473 |
apply (simp add: req_def knows.simps, safe) |
|
474 |
apply (erule in_synth_Guard, erule Guard_analz, simp) |
|
475 |
by (auto simp: req_def pro_def dest: Says_imp_knows_Spy) |
|
476 |
||
477 |
lemma req_imp_Guard_Friend: "[| evs:p1; A ~:bad; req A r n I B:set evs |] |
|
478 |
==> Guard n {priK A} (knows_max (Friend C) evs)" |
|
479 |
apply (rule Guard_knows_max') |
|
480 |
apply (rule_tac H="spies evs" in Guard_mono) |
|
481 |
apply (rule req_imp_Guard, simp+) |
|
482 |
apply (rule_tac B="spies' evs" in subset_trans) |
|
483 |
apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) |
|
484 |
by (rule knows'_sub_knows) |
|
485 |
||
486 |
subsubsection{*propositions are guarded*} |
|
487 |
||
488 |
lemma pro_imp_Guard [rule_format]: "[| evs:p1; B ~:bad; A ~:bad |] ==> |
|
489 |
pro B ofr A r I (cons M L) J C:set evs --> Guard ofr {priK A} (spies evs)" |
|
490 |
apply (erule p1.induct) (* +3 subgoals *) |
|
491 |
(* Nil *) |
|
492 |
apply simp |
|
493 |
(* Fake *) |
|
494 |
apply (simp add: pro_def, safe) (* +4 subgoals *) |
|
495 |
(* 1 *) |
|
496 |
apply (erule in_synth_Guard, drule Guard_analz, simp, simp) |
|
497 |
(* 2 *) |
|
498 |
apply simp |
|
499 |
(* 3 *) |
|
500 |
apply (simp, simp add: req_def pro_def, blast) |
|
501 |
(* 4 *) |
|
502 |
apply (simp add: pro_def) |
|
503 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard) |
|
504 |
(* 5 *) |
|
505 |
apply simp |
|
506 |
apply safe (* +1 subgoal *) |
|
507 |
apply (simp add: pro_def) |
|
508 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard) |
|
509 |
(* 6 *) |
|
510 |
apply (simp add: pro_def) |
|
511 |
apply (blast dest: Says_imp_knows_Spy) |
|
512 |
(* Request *) |
|
513 |
apply (simp add: pro_def) |
|
514 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard) |
|
515 |
(* Propose *) |
|
516 |
apply simp |
|
517 |
apply safe (* +1 subgoal *) |
|
518 |
(* 1 *) |
|
519 |
apply (simp add: pro_def) |
|
520 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard) |
|
521 |
(* 2 *) |
|
522 |
apply (simp add: pro_def) |
|
523 |
by (blast dest: Says_imp_knows_Spy) |
|
524 |
||
525 |
lemma pro_imp_Guard_Friend: "[| evs:p1; B ~:bad; A ~:bad; |
|
526 |
pro B ofr A r I (cons M L) J C:set evs |] |
|
527 |
==> Guard ofr {priK A} (knows_max (Friend D) evs)" |
|
528 |
apply (rule Guard_knows_max') |
|
529 |
apply (rule_tac H="spies evs" in Guard_mono) |
|
530 |
apply (rule pro_imp_Guard, simp+) |
|
531 |
apply (rule_tac B="spies' evs" in subset_trans) |
|
532 |
apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) |
|
533 |
by (rule knows'_sub_knows) |
|
534 |
||
535 |
subsubsection{*data confidentiality: |
|
536 |
no one other than the originator can decrypt the offers*} |
|
537 |
||
538 |
lemma Nonce_req_notin_spies: "[| evs:p1; req A r n I B:set evs; A ~:bad |] |
|
539 |
==> Nonce n ~:analz (spies evs)" |
|
540 |
by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) |
|
541 |
||
542 |
lemma Nonce_req_notin_knows_max_Friend: "[| evs:p1; req A r n I B:set evs; |
|
543 |
A ~:bad; A ~= Friend C |] ==> Nonce n ~:analz (knows_max (Friend C) evs)" |
|
544 |
apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+) |
|
545 |
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) |
|
546 |
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) |
|
547 |
||
548 |
lemma Nonce_pro_notin_spies: "[| evs:p1; B ~:bad; A ~:bad; |
|
549 |
pro B ofr A r I (cons M L) J C:set evs |] ==> Nonce ofr ~:analz (spies evs)" |
|
550 |
by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) |
|
551 |
||
552 |
lemma Nonce_pro_notin_knows_max_Friend: "[| evs:p1; B ~:bad; A ~:bad; |
|
553 |
A ~= Friend D; pro B ofr A r I (cons M L) J C:set evs |] |
|
554 |
==> Nonce ofr ~:analz (knows_max (Friend D) evs)" |
|
555 |
apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+) |
|
556 |
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) |
|
557 |
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) |
|
558 |
||
559 |
subsubsection{*non repudiability: |
|
560 |
an offer signed by B has been sent by B*} |
|
561 |
||
562 |
lemma Crypt_reqm: "[| Crypt (priK A) X:parts {reqm A' r n I B}; I:agl |] ==> A=A'" |
|
563 |
by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl) |
|
564 |
||
565 |
lemma Crypt_prom: "[| Crypt (priK A) X:parts {prom B ofr A' r I L J C}; |
|
566 |
I:agl; J:agl |] ==> A=B | Crypt (priK A) X:parts {L}" |
|
567 |
apply (simp add: prom_def anchor_def chain_def sign_def) |
|
568 |
by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel) |
|
569 |
||
570 |
lemma Crypt_safeness: "[| evs:p1; A ~:bad |] ==> Crypt (priK A) X:parts (spies evs) |
|
571 |
--> (EX B Y. Says A B Y:set evs & Crypt (priK A) X:parts {Y})" |
|
572 |
apply (erule p1.induct) |
|
573 |
(* Nil *) |
|
574 |
apply simp |
|
575 |
(* Fake *) |
|
576 |
apply clarsimp |
|
577 |
apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) |
|
578 |
apply (erule disjE) |
|
579 |
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) |
|
580 |
(* Request *) |
|
581 |
apply (simp add: req_def, clarify) |
|
582 |
apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) |
|
583 |
apply (erule disjE) |
|
584 |
apply (frule Crypt_reqm, simp, clarify) |
|
585 |
apply (rule_tac x=B in exI, rule_tac x="reqm A r n I B" in exI, simp, blast) |
|
586 |
(* Propose *) |
|
587 |
apply (simp add: pro_def, clarify) |
|
588 |
apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) |
|
589 |
apply (rotate_tac -1, erule disjE) |
|
590 |
apply (frule Crypt_prom, simp, simp) |
|
591 |
apply (rotate_tac -1, erule disjE) |
|
592 |
apply (rule_tac x=C in exI) |
|
593 |
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI, blast) |
|
594 |
apply (subgoal_tac "cons M L:parts (spies evsp)") |
|
595 |
apply (drule_tac G="{cons M L}" and H="spies evsp" in parts_trans, blast, blast) |
|
596 |
apply (drule Says_imp_spies, rotate_tac -1, drule parts.Inj) |
|
597 |
apply (drule parts.Snd, drule parts.Snd, drule parts.Snd) |
|
598 |
by auto |
|
599 |
||
600 |
lemma Crypt_Hash_imp_sign: "[| evs:p1; A ~:bad |] ==> |
|
601 |
Crypt (priK A) (Hash X):parts (spies evs) |
|
602 |
--> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})" |
|
603 |
apply (erule p1.induct) |
|
604 |
(* Nil *) |
|
605 |
apply simp |
|
606 |
(* Fake *) |
|
607 |
apply clarsimp |
|
608 |
apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) |
|
609 |
apply simp |
|
610 |
apply (erule disjE) |
|
611 |
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) |
|
612 |
(* Request *) |
|
613 |
apply (simp add: req_def, clarify) |
|
614 |
apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) |
|
615 |
apply simp |
|
616 |
apply (erule disjE) |
|
617 |
apply (frule Crypt_reqm, simp+) |
|
618 |
apply (rule_tac x=B in exI, rule_tac x="reqm Aa r n I B" in exI) |
|
619 |
apply (simp add: reqm_def sign_def anchor_def no_Crypt_in_agl) |
|
620 |
apply (simp add: chain_def sign_def, blast) |
|
621 |
(* Propose *) |
|
622 |
apply (simp add: pro_def, clarify) |
|
623 |
apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) |
|
624 |
apply simp |
|
625 |
apply (rotate_tac -1, erule disjE) |
|
626 |
apply (simp add: prom_def sign_def no_Crypt_in_agl no_Crypt_in_appdel) |
|
627 |
apply (simp add: chain_def sign_def) |
|
628 |
apply (rotate_tac -1, erule disjE) |
|
629 |
apply (rule_tac x=C in exI) |
|
630 |
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI) |
|
631 |
apply (simp add: prom_def chain_def sign_def) |
|
632 |
apply (erule impE) |
|
633 |
apply (blast dest: get_ML parts_sub) |
|
634 |
apply (blast del: MPair_parts)+ |
|
635 |
done |
|
636 |
||
637 |
lemma sign_safeness: "[| evs:p1; A ~:bad |] ==> sign A X:parts (spies evs) |
|
638 |
--> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})" |
|
639 |
apply (clarify, simp add: sign_def, frule parts.Snd) |
|
640 |
apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def]) |
|
641 |
done |
|
642 |
||
643 |
end |