| author | wenzelm | 
| Sun, 29 Jul 2007 17:28:56 +0200 | |
| changeset 24059 | 89a5382406a1 | 
| parent 23746 | a455e69c31cc | 
| child 32960 | 69916a850301 | 
| permissions | -rwxr-xr-x | 
| 18886 | 1  | 
(* ID: $Id$  | 
2  | 
Author: Giampaolo Bella, Catania University  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header{*The Kerberos Protocol, BAN Version, with Gets event*}
 | 
|
6  | 
||
7  | 
theory Kerberos_BAN_Gets imports Public begin  | 
|
8  | 
||
9  | 
text{*From page 251 of
 | 
|
10  | 
Burrows, Abadi and Needham (1989). A Logic of Authentication.  | 
|
11  | 
Proc. Royal Soc. 426  | 
|
12  | 
||
13  | 
Confidentiality (secrecy) and authentication properties rely on  | 
|
14  | 
temporal checks: strong guarantees in a little abstracted - but  | 
|
15  | 
very realistic - model.  | 
|
16  | 
*}  | 
|
17  | 
||
18  | 
(* Temporal modelization: session keys can be leaked  | 
|
19  | 
ONLY when they have expired *)  | 
|
20  | 
||
21  | 
consts  | 
|
22  | 
||
23  | 
(*Duration of the session key*)  | 
|
24  | 
sesKlife :: nat  | 
|
25  | 
||
26  | 
(*Duration of the authenticator*)  | 
|
27  | 
authlife :: nat  | 
|
28  | 
||
29  | 
text{*The ticket should remain fresh for two journeys on the network at least*}
 | 
|
30  | 
text{*The Gets event causes longer traces for the protocol to reach its end*}
 | 
|
31  | 
specification (sesKlife)  | 
|
32  | 
sesKlife_LB [iff]: "4 \<le> sesKlife"  | 
|
33  | 
by blast  | 
|
34  | 
||
35  | 
text{*The authenticator only for one journey*}
 | 
|
36  | 
text{*The Gets event causes longer traces for the protocol to reach its end*}
 | 
|
37  | 
specification (authlife)  | 
|
38  | 
authlife_LB [iff]: "2 \<le> authlife"  | 
|
39  | 
by blast  | 
|
40  | 
||
41  | 
||
| 20768 | 42  | 
abbreviation  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20768 
diff
changeset
 | 
43  | 
CT :: "event list=>nat" where  | 
| 20768 | 44  | 
"CT == length"  | 
| 18886 | 45  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20768 
diff
changeset
 | 
46  | 
abbreviation  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20768 
diff
changeset
 | 
47  | 
expiredK :: "[nat, event list] => bool" where  | 
| 20768 | 48  | 
"expiredK T evs == sesKlife + T < CT evs"  | 
| 18886 | 49  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20768 
diff
changeset
 | 
50  | 
abbreviation  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20768 
diff
changeset
 | 
51  | 
expiredA :: "[nat, event list] => bool" where  | 
| 20768 | 52  | 
"expiredA T evs == authlife + T < CT evs"  | 
| 18886 | 53  | 
|
54  | 
||
55  | 
constdefs  | 
|
56  | 
(* Yields the subtrace of a given trace from its beginning to a given event *)  | 
|
57  | 
  before :: "[event, event list] => event list" ("before _ on _")
 | 
|
58  | 
"before ev on evs == takeWhile (% z. z ~= ev) (rev evs)"  | 
|
59  | 
||
60  | 
(* States than an event really appears only once on a trace *)  | 
|
61  | 
  Unique :: "[event, event list] => bool" ("Unique _ on _")
 | 
|
62  | 
"Unique ev on evs ==  | 
|
63  | 
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"  | 
|
64  | 
||
65  | 
||
| 23746 | 66  | 
inductive_set bankerb_gets :: "event list set"  | 
67  | 
where  | 
|
| 18886 | 68  | 
|
69  | 
Nil: "[] \<in> bankerb_gets"  | 
|
70  | 
||
| 23746 | 71  | 
| Fake: "\<lbrakk> evsf \<in> bankerb_gets; X \<in> synth (analz (knows Spy evsf)) \<rbrakk>  | 
| 18886 | 72  | 
\<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"  | 
73  | 
||
| 23746 | 74  | 
| Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>  | 
| 18886 | 75  | 
\<Longrightarrow> Gets B X # evsr \<in> bankerb_gets"  | 
76  | 
||
| 23746 | 77  | 
| BK1: "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>  | 
| 18886 | 78  | 
\<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1  | 
79  | 
\<in> bankerb_gets"  | 
|
80  | 
||
81  | 
||
| 23746 | 82  | 
| BK2: "\<lbrakk> evs2 \<in> bankerb_gets; Key K \<notin> used evs2; K \<in> symKeys;  | 
| 18886 | 83  | 
Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>  | 
84  | 
\<Longrightarrow> Says Server A  | 
|
85  | 
(Crypt (shrK A)  | 
|
86  | 
\<lbrace>Number (CT evs2), Agent B, Key K,  | 
|
87  | 
(Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)  | 
|
88  | 
# evs2 \<in> bankerb_gets"  | 
|
89  | 
||
90  | 
||
| 23746 | 91  | 
| BK3: "\<lbrakk> evs3 \<in> bankerb_gets;  | 
| 18886 | 92  | 
Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)  | 
93  | 
\<in> set evs3;  | 
|
94  | 
Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;  | 
|
95  | 
\<not> expiredK Tk evs3 \<rbrakk>  | 
|
96  | 
\<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>  | 
|
97  | 
# evs3 \<in> bankerb_gets"  | 
|
98  | 
||
99  | 
||
| 23746 | 100  | 
| BK4: "\<lbrakk> evs4 \<in> bankerb_gets;  | 
| 18886 | 101  | 
Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),  | 
102  | 
(Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;  | 
|
103  | 
\<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>  | 
|
104  | 
\<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4  | 
|
105  | 
\<in> bankerb_gets"  | 
|
106  | 
||
107  | 
(*Old session keys may become compromised*)  | 
|
| 23746 | 108  | 
| Oops: "\<lbrakk> evso \<in> bankerb_gets;  | 
| 18886 | 109  | 
Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)  | 
110  | 
\<in> set evso;  | 
|
111  | 
expiredK Tk evso \<rbrakk>  | 
|
112  | 
\<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"  | 
|
113  | 
||
114  | 
||
115  | 
declare Says_imp_knows_Spy [THEN parts.Inj, dest]  | 
|
116  | 
declare parts.Body [dest]  | 
|
117  | 
declare analz_into_parts [dest]  | 
|
118  | 
declare Fake_parts_insert_in_Un [dest]  | 
|
119  | 
declare knows_Spy_partsEs [elim]  | 
|
120  | 
||
121  | 
||
122  | 
text{*A "possibility property": there are traces that reach the end.*}
 | 
|
123  | 
lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk>  | 
|
124  | 
\<Longrightarrow> \<exists>Timestamp. \<exists>evs \<in> bankerb_gets.  | 
|
125  | 
Says B A (Crypt K (Number Timestamp))  | 
|
126  | 
\<in> set evs"  | 
|
127  | 
apply (cut_tac sesKlife_LB)  | 
|
128  | 
apply (cut_tac authlife_LB)  | 
|
129  | 
apply (intro exI bexI)  | 
|
130  | 
apply (rule_tac [2]  | 
|
131  | 
bankerb_gets.Nil [THEN bankerb_gets.BK1, THEN bankerb_gets.Reception,  | 
|
132  | 
THEN bankerb_gets.BK2, THEN bankerb_gets.Reception,  | 
|
133  | 
THEN bankerb_gets.BK3, THEN bankerb_gets.Reception,  | 
|
134  | 
THEN bankerb_gets.BK4])  | 
|
135  | 
apply (possibility, simp_all (no_asm_simp) add: used_Cons)  | 
|
136  | 
done  | 
|
137  | 
||
138  | 
||
139  | 
text{*Lemmas about reception invariant: if a message is received it certainly
 | 
|
140  | 
was sent*}  | 
|
141  | 
lemma Gets_imp_Says :  | 
|
142  | 
"\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"  | 
|
143  | 
apply (erule rev_mp)  | 
|
144  | 
apply (erule bankerb_gets.induct)  | 
|
145  | 
apply auto  | 
|
146  | 
done  | 
|
147  | 
||
148  | 
lemma Gets_imp_knows_Spy:  | 
|
149  | 
"\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> X \<in> knows Spy evs"  | 
|
150  | 
apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy)  | 
|
151  | 
done  | 
|
152  | 
||
153  | 
lemma Gets_imp_knows_Spy_parts[dest]:  | 
|
154  | 
"\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> X \<in> parts (knows Spy evs)"  | 
|
155  | 
apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj])  | 
|
156  | 
done  | 
|
157  | 
||
158  | 
lemma Gets_imp_knows:  | 
|
159  | 
"\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> X \<in> knows B evs"  | 
|
160  | 
apply (case_tac "B = Spy")  | 
|
161  | 
apply (blast dest!: Gets_imp_knows_Spy)  | 
|
162  | 
apply (blast dest!: Gets_imp_knows_agents)  | 
|
163  | 
done  | 
|
164  | 
||
165  | 
lemma Gets_imp_knows_analz:  | 
|
166  | 
"\<lbrakk> Gets B X \<in> set evs; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> X \<in> analz (knows B evs)"  | 
|
167  | 
apply (blast dest: Gets_imp_knows [THEN analz.Inj])  | 
|
168  | 
done  | 
|
169  | 
||
170  | 
text{*Lemmas for reasoning about predicate "before"*}
 | 
|
171  | 
lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)";
 | 
|
172  | 
apply (induct_tac "evs")  | 
|
173  | 
apply simp  | 
|
174  | 
apply (induct_tac "a")  | 
|
175  | 
apply auto  | 
|
176  | 
done  | 
|
177  | 
||
178  | 
lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)";
 | 
|
179  | 
apply (induct_tac "evs")  | 
|
180  | 
apply simp  | 
|
181  | 
apply (induct_tac "a")  | 
|
182  | 
apply auto  | 
|
183  | 
done  | 
|
184  | 
||
185  | 
lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs";  | 
|
186  | 
apply (induct_tac "evs")  | 
|
187  | 
apply simp  | 
|
188  | 
apply (induct_tac "a")  | 
|
189  | 
apply auto  | 
|
190  | 
done  | 
|
191  | 
||
192  | 
lemma used_evs_rev: "used evs = used (rev evs)"  | 
|
193  | 
apply (induct_tac "evs")  | 
|
194  | 
apply simp  | 
|
195  | 
apply (induct_tac "a")  | 
|
196  | 
apply (simp add: used_Says_rev)  | 
|
197  | 
apply (simp add: used_Gets_rev)  | 
|
198  | 
apply (simp add: used_Notes_rev)  | 
|
199  | 
done  | 
|
200  | 
||
201  | 
lemma used_takeWhile_used [rule_format]:  | 
|
202  | 
"x : used (takeWhile P X) --> x : used X"  | 
|
203  | 
apply (induct_tac "X")  | 
|
204  | 
apply simp  | 
|
205  | 
apply (induct_tac "a")  | 
|
206  | 
apply (simp_all add: used_Nil)  | 
|
207  | 
apply (blast dest!: initState_into_used)+  | 
|
208  | 
done  | 
|
209  | 
||
210  | 
lemma set_evs_rev: "set evs = set (rev evs)"  | 
|
211  | 
apply auto  | 
|
212  | 
done  | 
|
213  | 
||
214  | 
lemma takeWhile_void [rule_format]:  | 
|
215  | 
"x \<notin> set evs \<longrightarrow> takeWhile (\<lambda>z. z \<noteq> x) evs = evs"  | 
|
216  | 
apply auto  | 
|
217  | 
done  | 
|
218  | 
||
219  | 
(**** Inductive proofs about bankerb_gets ****)  | 
|
220  | 
||
221  | 
text{*Forwarding Lemma for reasoning about the encrypted portion of message BK3*}
 | 
|
222  | 
lemma BK3_msg_in_parts_knows_Spy:  | 
|
223  | 
"\<lbrakk>Gets A (Crypt KA \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs; evs \<in> bankerb_gets \<rbrakk>  | 
|
224  | 
\<Longrightarrow> X \<in> parts (knows Spy evs)"  | 
|
225  | 
apply blast  | 
|
226  | 
done  | 
|
227  | 
||
228  | 
lemma Oops_parts_knows_Spy:  | 
|
229  | 
"Says Server A (Crypt (shrK A) \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs  | 
|
230  | 
\<Longrightarrow> K \<in> parts (knows Spy evs)"  | 
|
231  | 
apply blast  | 
|
232  | 
done  | 
|
233  | 
||
234  | 
||
235  | 
text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
 | 
|
236  | 
lemma Spy_see_shrK [simp]:  | 
|
237  | 
"evs \<in> bankerb_gets \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"  | 
|
238  | 
apply (erule bankerb_gets.induct)  | 
|
239  | 
apply (frule_tac [8] Oops_parts_knows_Spy)  | 
|
240  | 
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast+)  | 
|
241  | 
done  | 
|
242  | 
||
243  | 
||
244  | 
lemma Spy_analz_shrK [simp]:  | 
|
245  | 
"evs \<in> bankerb_gets \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"  | 
|
246  | 
by auto  | 
|
247  | 
||
248  | 
lemma Spy_see_shrK_D [dest!]:  | 
|
249  | 
"\<lbrakk> Key (shrK A) \<in> parts (knows Spy evs);  | 
|
250  | 
evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> A:bad"  | 
|
251  | 
by (blast dest: Spy_see_shrK)  | 
|
252  | 
||
253  | 
lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]  | 
|
254  | 
||
255  | 
||
256  | 
text{*Nobody can have used non-existent keys!*}
 | 
|
257  | 
lemma new_keys_not_used [simp]:  | 
|
258  | 
"\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> bankerb_gets\<rbrakk>  | 
|
259  | 
\<Longrightarrow> K \<notin> keysFor (parts (knows Spy evs))"  | 
|
260  | 
apply (erule rev_mp)  | 
|
261  | 
apply (erule bankerb_gets.induct)  | 
|
262  | 
apply (frule_tac [8] Oops_parts_knows_Spy)  | 
|
263  | 
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)  | 
|
264  | 
txt{*Fake*}
 | 
|
265  | 
apply (force dest!: keysFor_parts_insert)  | 
|
266  | 
txt{*BK2, BK3, BK4*}
 | 
|
267  | 
apply (force dest!: analz_shrK_Decrypt)+  | 
|
268  | 
done  | 
|
269  | 
||
270  | 
subsection{* Lemmas concerning the form of items passed in messages *}
 | 
|
271  | 
||
272  | 
text{*Describes the form of K, X and K' when the Server sends this message.*}
 | 
|
273  | 
lemma Says_Server_message_form:  | 
|
274  | 
"\<lbrakk> Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)  | 
|
275  | 
\<in> set evs; evs \<in> bankerb_gets \<rbrakk>  | 
|
276  | 
\<Longrightarrow> K' = shrK A & K \<notin> range shrK &  | 
|
277  | 
Ticket = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>) &  | 
|
278  | 
Key K \<notin> used(before  | 
|
279  | 
Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)  | 
|
280  | 
on evs) &  | 
|
281  | 
Tk = CT(before  | 
|
282  | 
Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)  | 
|
283  | 
on evs)"  | 
|
284  | 
apply (unfold before_def)  | 
|
285  | 
apply (erule rev_mp)  | 
|
286  | 
apply (erule bankerb_gets.induct, simp_all)  | 
|
287  | 
txt{*We need this simplification only for Message 2*}
 | 
|
288  | 
apply (simp (no_asm) add: takeWhile_tail)  | 
|
289  | 
apply auto  | 
|
290  | 
txt{*Two subcases of Message 2. Subcase: used before*}
 | 
|
291  | 
apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD]  | 
|
292  | 
used_takeWhile_used)  | 
|
293  | 
txt{*subcase: CT before*}
 | 
|
294  | 
apply (fastsimp dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])  | 
|
295  | 
done  | 
|
296  | 
||
297  | 
||
298  | 
text{*If the encrypted message appears then it originated with the Server
 | 
|
299  | 
PROVIDED that A is NOT compromised!  | 
|
300  | 
This allows A to verify freshness of the session key.  | 
|
301  | 
*}  | 
|
302  | 
lemma Kab_authentic:  | 
|
303  | 
"\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>  | 
|
304  | 
\<in> parts (knows Spy evs);  | 
|
305  | 
A \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
306  | 
\<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)  | 
|
307  | 
\<in> set evs"  | 
|
308  | 
apply (erule rev_mp)  | 
|
309  | 
apply (erule bankerb_gets.induct)  | 
|
310  | 
apply (frule_tac [8] Oops_parts_knows_Spy)  | 
|
311  | 
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast)  | 
|
312  | 
done  | 
|
313  | 
||
314  | 
||
315  | 
text{*If the TICKET appears then it originated with the Server*}
 | 
|
316  | 
text{*FRESHNESS OF THE SESSION KEY to B*}
 | 
|
317  | 
lemma ticket_authentic:  | 
|
318  | 
"\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (knows Spy evs);  | 
|
319  | 
B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
320  | 
\<Longrightarrow> Says Server A  | 
|
321  | 
(Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,  | 
|
322  | 
Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)  | 
|
323  | 
\<in> set evs"  | 
|
324  | 
apply (erule rev_mp)  | 
|
325  | 
apply (erule bankerb_gets.induct)  | 
|
326  | 
apply (frule_tac [8] Oops_parts_knows_Spy)  | 
|
327  | 
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast)  | 
|
328  | 
done  | 
|
329  | 
||
330  | 
||
331  | 
text{*EITHER describes the form of X when the following message is sent,
 | 
|
332  | 
OR reduces it to the Fake case.  | 
|
333  | 
  Use @{text Says_Server_message_form} if applicable.*}
 | 
|
334  | 
lemma Gets_Server_message_form:  | 
|
335  | 
"\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)  | 
|
336  | 
\<in> set evs;  | 
|
337  | 
evs \<in> bankerb_gets \<rbrakk>  | 
|
338  | 
\<Longrightarrow> (K \<notin> range shrK & X = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>))  | 
|
339  | 
| X \<in> analz (knows Spy evs)"  | 
|
340  | 
apply (case_tac "A \<in> bad")  | 
|
341  | 
apply (force dest!: Gets_imp_knows_Spy [THEN analz.Inj])  | 
|
342  | 
apply (blast dest!: Kab_authentic Says_Server_message_form)  | 
|
343  | 
done  | 
|
344  | 
||
345  | 
||
346  | 
text{*Reliability guarantees: honest agents act as we expect*}
 | 
|
347  | 
||
348  | 
lemma BK3_imp_Gets:  | 
|
349  | 
"\<lbrakk> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs;  | 
|
350  | 
A \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
351  | 
\<Longrightarrow> \<exists> Tk. Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)  | 
|
352  | 
\<in> set evs"  | 
|
353  | 
apply (erule rev_mp)  | 
|
354  | 
apply (erule bankerb_gets.induct)  | 
|
355  | 
apply auto  | 
|
356  | 
done  | 
|
357  | 
||
358  | 
lemma BK4_imp_Gets:  | 
|
359  | 
"\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;  | 
|
360  | 
B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
361  | 
\<Longrightarrow> \<exists> Tk. Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,  | 
|
362  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"  | 
|
363  | 
apply (erule rev_mp)  | 
|
364  | 
apply (erule bankerb_gets.induct)  | 
|
365  | 
apply auto  | 
|
366  | 
done  | 
|
367  | 
||
368  | 
lemma Gets_A_knows_K:  | 
|
369  | 
"\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;  | 
|
370  | 
evs \<in> bankerb_gets \<rbrakk>  | 
|
371  | 
\<Longrightarrow> Key K \<in> analz (knows A evs)"  | 
|
372  | 
apply (force dest: Gets_imp_knows_analz)  | 
|
373  | 
done  | 
|
374  | 
||
375  | 
lemma Gets_B_knows_K:  | 
|
376  | 
"\<lbrakk> Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,  | 
|
377  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs;  | 
|
378  | 
evs \<in> bankerb_gets \<rbrakk>  | 
|
379  | 
\<Longrightarrow> Key K \<in> analz (knows B evs)"  | 
|
380  | 
apply (force dest: Gets_imp_knows_analz)  | 
|
381  | 
done  | 
|
382  | 
||
383  | 
||
384  | 
(****  | 
|
385  | 
The following is to prove theorems of the form  | 
|
386  | 
||
387  | 
Key K \<in> analz (insert (Key KAB) (knows Spy evs)) \<Longrightarrow>  | 
|
388  | 
Key K \<in> analz (knows Spy evs)  | 
|
389  | 
||
390  | 
A more general formula must be proved inductively.  | 
|
391  | 
||
392  | 
****)  | 
|
393  | 
||
394  | 
||
395  | 
text{* Session keys are not used to encrypt other session keys *}
 | 
|
396  | 
lemma analz_image_freshK [rule_format (no_asm)]:  | 
|
397  | 
"evs \<in> bankerb_gets \<Longrightarrow>  | 
|
398  | 
\<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>  | 
|
399  | 
(Key K \<in> analz (Key`KK Un (knows Spy evs))) =  | 
|
400  | 
(K \<in> KK | Key K \<in> analz (knows Spy evs))"  | 
|
401  | 
apply (erule bankerb_gets.induct)  | 
|
402  | 
apply (drule_tac [8] Says_Server_message_form)  | 
|
403  | 
apply (erule_tac [6] Gets_Server_message_form [THEN disjE], analz_freshK, spy_analz, auto)  | 
|
404  | 
done  | 
|
405  | 
||
406  | 
||
407  | 
lemma analz_insert_freshK:  | 
|
408  | 
"\<lbrakk> evs \<in> bankerb_gets; KAB \<notin> range shrK \<rbrakk> \<Longrightarrow>  | 
|
409  | 
(Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =  | 
|
410  | 
(K = KAB | Key K \<in> analz (knows Spy evs))"  | 
|
411  | 
by (simp only: analz_image_freshK analz_image_freshK_simps)  | 
|
412  | 
||
413  | 
||
414  | 
text{* The session key K uniquely identifies the message *}
 | 
|
415  | 
lemma unique_session_keys:  | 
|
416  | 
"\<lbrakk> Says Server A  | 
|
417  | 
(Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;  | 
|
418  | 
Says Server A'  | 
|
419  | 
(Crypt (shrK A') \<lbrace>Number Tk', Agent B', Key K, X'\<rbrace>) \<in> set evs;  | 
|
420  | 
evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> A=A' & Tk=Tk' & B=B' & X = X'"  | 
|
421  | 
apply (erule rev_mp)  | 
|
422  | 
apply (erule rev_mp)  | 
|
423  | 
apply (erule bankerb_gets.induct)  | 
|
424  | 
apply (frule_tac [8] Oops_parts_knows_Spy)  | 
|
425  | 
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)  | 
|
426  | 
txt{*BK2: it can't be a new key*}
 | 
|
427  | 
apply blast  | 
|
428  | 
done  | 
|
429  | 
||
430  | 
lemma unique_session_keys_Gets:  | 
|
431  | 
"\<lbrakk> Gets A  | 
|
432  | 
(Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;  | 
|
433  | 
Gets A  | 
|
434  | 
(Crypt (shrK A) \<lbrace>Number Tk', Agent B', Key K, X'\<rbrace>) \<in> set evs;  | 
|
435  | 
A \<notin> bad; evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow> Tk=Tk' & B=B' & X = X'"  | 
|
436  | 
apply (blast dest!: Kab_authentic unique_session_keys)  | 
|
437  | 
done  | 
|
438  | 
||
439  | 
||
440  | 
lemma Server_Unique:  | 
|
441  | 
"\<lbrakk> Says Server A  | 
|
442  | 
(Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;  | 
|
443  | 
evs \<in> bankerb_gets \<rbrakk> \<Longrightarrow>  | 
|
444  | 
Unique Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)  | 
|
445  | 
on evs"  | 
|
446  | 
apply (erule rev_mp, erule bankerb_gets.induct, simp_all add: Unique_def)  | 
|
447  | 
apply blast  | 
|
448  | 
done  | 
|
449  | 
||
450  | 
||
451  | 
||
452  | 
subsection{*Non-temporal guarantees, explicitly relying on non-occurrence of
 | 
|
453  | 
oops events - refined below by temporal guarantees*}  | 
|
454  | 
||
455  | 
text{*Non temporal treatment of confidentiality*}
 | 
|
456  | 
||
457  | 
text{* Lemma: the session key sent in msg BK2 would be lost by oops
 | 
|
458  | 
if the spy could see it! *}  | 
|
459  | 
lemma lemma_conf [rule_format (no_asm)]:  | 
|
460  | 
"\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
461  | 
\<Longrightarrow> Says Server A  | 
|
462  | 
(Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,  | 
|
463  | 
Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)  | 
|
464  | 
\<in> set evs \<longrightarrow>  | 
|
465  | 
Key K \<in> analz (knows Spy evs) \<longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<in> set evs"  | 
|
466  | 
apply (erule bankerb_gets.induct)  | 
|
467  | 
apply (frule_tac [8] Says_Server_message_form)  | 
|
468  | 
apply (frule_tac [6] Gets_Server_message_form [THEN disjE])  | 
|
469  | 
apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes)  | 
|
470  | 
txt{*Fake*}
 | 
|
471  | 
apply spy_analz  | 
|
472  | 
txt{*BK2*}
 | 
|
473  | 
apply (blast intro: parts_insertI)  | 
|
474  | 
txt{*BK3*}
 | 
|
475  | 
apply (case_tac "Aa \<in> bad")  | 
|
476  | 
prefer 2 apply (blast dest: Kab_authentic unique_session_keys)  | 
|
477  | 
apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz)  | 
|
478  | 
txt{*Oops*}
 | 
|
479  | 
apply (blast dest: unique_session_keys)  | 
|
480  | 
done  | 
|
481  | 
||
482  | 
||
483  | 
text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
 | 
|
484  | 
as long as they have not expired.*}  | 
|
485  | 
lemma Confidentiality_S:  | 
|
486  | 
"\<lbrakk> Says Server A  | 
|
487  | 
(Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;  | 
|
488  | 
Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;  | 
|
489  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets  | 
|
490  | 
\<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"  | 
|
491  | 
apply (frule Says_Server_message_form, assumption)  | 
|
492  | 
apply (blast intro: lemma_conf)  | 
|
493  | 
done  | 
|
494  | 
||
495  | 
text{*Confidentiality for Alice*}
 | 
|
496  | 
lemma Confidentiality_A:  | 
|
497  | 
"\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);  | 
|
498  | 
Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;  | 
|
499  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets  | 
|
500  | 
\<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"  | 
|
501  | 
by (blast dest!: Kab_authentic Confidentiality_S)  | 
|
502  | 
||
503  | 
text{*Confidentiality for Bob*}
 | 
|
504  | 
lemma Confidentiality_B:  | 
|
505  | 
"\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  | 
|
506  | 
\<in> parts (knows Spy evs);  | 
|
507  | 
Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;  | 
|
508  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets  | 
|
509  | 
\<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"  | 
|
510  | 
by (blast dest!: ticket_authentic Confidentiality_S)  | 
|
511  | 
||
512  | 
||
513  | 
text{*Non temporal treatment of authentication*}
 | 
|
514  | 
||
515  | 
text{*Lemmas @{text lemma_A} and @{text lemma_B} in fact are common to both temporal and non-temporal treatments*}
 | 
|
516  | 
lemma lemma_A [rule_format]:  | 
|
517  | 
"\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
518  | 
\<Longrightarrow>  | 
|
519  | 
Key K \<notin> analz (knows Spy evs) \<longrightarrow>  | 
|
520  | 
Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)  | 
|
521  | 
\<in> set evs \<longrightarrow>  | 
|
522  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>  | 
|
523  | 
Says A B \<lbrace>X, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace>  | 
|
524  | 
\<in> set evs"  | 
|
525  | 
apply (erule bankerb_gets.induct)  | 
|
526  | 
apply (frule_tac [8] Oops_parts_knows_Spy)  | 
|
527  | 
apply (frule_tac [6] Gets_Server_message_form)  | 
|
528  | 
apply (frule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)  | 
|
529  | 
apply (simp_all (no_asm_simp) add: all_conj_distrib)  | 
|
530  | 
txt{*Fake*}
 | 
|
531  | 
apply blast  | 
|
532  | 
txt{*BK2*}
 | 
|
533  | 
apply (force dest: Crypt_imp_invKey_keysFor)  | 
|
534  | 
txt{*BK3*}
 | 
|
535  | 
apply (blast dest: Kab_authentic unique_session_keys)  | 
|
536  | 
done  | 
|
537  | 
lemma lemma_B [rule_format]:  | 
|
538  | 
"\<lbrakk> B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
539  | 
\<Longrightarrow> Key K \<notin> analz (knows Spy evs) \<longrightarrow>  | 
|
540  | 
Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)  | 
|
541  | 
\<in> set evs \<longrightarrow>  | 
|
542  | 
Crypt K (Number Ta) \<in> parts (knows Spy evs) \<longrightarrow>  | 
|
543  | 
Says B A (Crypt K (Number Ta)) \<in> set evs"  | 
|
544  | 
apply (erule bankerb_gets.induct)  | 
|
545  | 
apply (frule_tac [8] Oops_parts_knows_Spy)  | 
|
546  | 
apply (frule_tac [6] Gets_Server_message_form)  | 
|
547  | 
apply (drule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)  | 
|
548  | 
apply (simp_all (no_asm_simp) add: all_conj_distrib)  | 
|
549  | 
txt{*Fake*}
 | 
|
550  | 
apply blast  | 
|
551  | 
txt{*BK2*} 
 | 
|
552  | 
apply (force dest: Crypt_imp_invKey_keysFor)  | 
|
553  | 
txt{*BK4*}
 | 
|
554  | 
apply (blast dest: ticket_authentic unique_session_keys  | 
|
555  | 
Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad)  | 
|
556  | 
done  | 
|
557  | 
||
558  | 
||
559  | 
text{*The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
 | 
|
560  | 
||
561  | 
text{*Authentication of A to B*}
 | 
|
562  | 
lemma B_authenticates_A_r:  | 
|
563  | 
"\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs);  | 
|
564  | 
Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (knows Spy evs);  | 
|
565  | 
Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;  | 
|
566  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
567  | 
\<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,  | 
|
568  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"  | 
|
569  | 
by (blast dest!: ticket_authentic  | 
|
570  | 
intro!: lemma_A  | 
|
571  | 
elim!: Confidentiality_S [THEN [2] rev_notE])  | 
|
572  | 
||
573  | 
text{*Authentication of B to A*}
 | 
|
574  | 
lemma A_authenticates_B_r:  | 
|
575  | 
"\<lbrakk> Crypt K (Number Ta) \<in> parts (knows Spy evs);  | 
|
576  | 
Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);  | 
|
577  | 
Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;  | 
|
578  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
579  | 
\<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"  | 
|
580  | 
by (blast dest!: Kab_authentic  | 
|
581  | 
intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE])  | 
|
582  | 
||
583  | 
lemma B_authenticates_A:  | 
|
584  | 
"\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);  | 
|
585  | 
Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs);  | 
|
586  | 
Key K \<notin> analz (spies evs);  | 
|
587  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
588  | 
\<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,  | 
|
589  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"  | 
|
590  | 
apply (blast dest!: ticket_authentic intro!: lemma_A)  | 
|
591  | 
done  | 
|
592  | 
||
593  | 
lemma A_authenticates_B:  | 
|
594  | 
"\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);  | 
|
595  | 
Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);  | 
|
596  | 
Key K \<notin> analz (spies evs);  | 
|
597  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
598  | 
\<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"  | 
|
599  | 
apply (blast dest!: Kab_authentic intro!: lemma_B)  | 
|
600  | 
done  | 
|
601  | 
||
602  | 
||
603  | 
subsection{*Temporal guarantees, relying on a temporal check that insures that
 | 
|
604  | 
no oops event occurred. These are available in the sense of goal availability*}  | 
|
605  | 
||
606  | 
||
607  | 
text{*Temporal treatment of confidentiality*}
 | 
|
608  | 
||
609  | 
text{* Lemma: the session key sent in msg BK2 would be EXPIRED
 | 
|
610  | 
if the spy could see it! *}  | 
|
611  | 
lemma lemma_conf_temporal [rule_format (no_asm)]:  | 
|
612  | 
"\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
613  | 
\<Longrightarrow> Says Server A  | 
|
614  | 
(Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,  | 
|
615  | 
Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)  | 
|
616  | 
\<in> set evs \<longrightarrow>  | 
|
617  | 
Key K \<in> analz (knows Spy evs) \<longrightarrow> expiredK Tk evs"  | 
|
618  | 
apply (erule bankerb_gets.induct)  | 
|
619  | 
apply (frule_tac [8] Says_Server_message_form)  | 
|
620  | 
apply (frule_tac [6] Gets_Server_message_form [THEN disjE])  | 
|
621  | 
apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)  | 
|
622  | 
txt{*Fake*}
 | 
|
623  | 
apply spy_analz  | 
|
624  | 
txt{*BK2*}
 | 
|
625  | 
apply (blast intro: parts_insertI less_SucI)  | 
|
626  | 
txt{*BK3*}
 | 
|
627  | 
apply (case_tac "Aa \<in> bad")  | 
|
628  | 
prefer 2 apply (blast dest: Kab_authentic unique_session_keys)  | 
|
629  | 
apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI)  | 
|
630  | 
txt{*Oops: PROOF FAILS if unsafe intro below*}
 | 
|
631  | 
apply (blast dest: unique_session_keys intro!: less_SucI)  | 
|
632  | 
done  | 
|
633  | 
||
634  | 
||
635  | 
text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
 | 
|
636  | 
as long as they have not expired.*}  | 
|
637  | 
lemma Confidentiality_S_temporal:  | 
|
638  | 
"\<lbrakk> Says Server A  | 
|
639  | 
(Crypt K' \<lbrace>Number T, Agent B, Key K, X\<rbrace>) \<in> set evs;  | 
|
640  | 
\<not> expiredK T evs;  | 
|
641  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets  | 
|
642  | 
\<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"  | 
|
643  | 
apply (frule Says_Server_message_form, assumption)  | 
|
644  | 
apply (blast intro: lemma_conf_temporal)  | 
|
645  | 
done  | 
|
646  | 
||
647  | 
text{*Confidentiality for Alice*}
 | 
|
648  | 
lemma Confidentiality_A_temporal:  | 
|
649  | 
"\<lbrakk> Crypt (shrK A) \<lbrace>Number T, Agent B, Key K, X\<rbrace> \<in> parts (knows Spy evs);  | 
|
650  | 
\<not> expiredK T evs;  | 
|
651  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets  | 
|
652  | 
\<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"  | 
|
653  | 
by (blast dest!: Kab_authentic Confidentiality_S_temporal)  | 
|
654  | 
||
655  | 
text{*Confidentiality for Bob*}
 | 
|
656  | 
lemma Confidentiality_B_temporal:  | 
|
657  | 
"\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  | 
|
658  | 
\<in> parts (knows Spy evs);  | 
|
659  | 
\<not> expiredK Tk evs;  | 
|
660  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets  | 
|
661  | 
\<rbrakk> \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"  | 
|
662  | 
by (blast dest!: ticket_authentic Confidentiality_S_temporal)  | 
|
663  | 
||
664  | 
||
665  | 
text{*Temporal treatment of authentication*}
 | 
|
666  | 
||
667  | 
text{*Authentication of A to B*}
 | 
|
668  | 
lemma B_authenticates_A_temporal:  | 
|
669  | 
"\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (knows Spy evs);  | 
|
670  | 
Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  | 
|
671  | 
\<in> parts (knows Spy evs);  | 
|
672  | 
\<not> expiredK Tk evs;  | 
|
673  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
674  | 
\<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,  | 
|
675  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"  | 
|
676  | 
by (blast dest!: ticket_authentic  | 
|
677  | 
intro!: lemma_A  | 
|
678  | 
elim!: Confidentiality_S_temporal [THEN [2] rev_notE])  | 
|
679  | 
||
680  | 
text{*Authentication of B to A*}
 | 
|
681  | 
lemma A_authenticates_B_temporal:  | 
|
682  | 
"\<lbrakk> Crypt K (Number Ta) \<in> parts (knows Spy evs);  | 
|
683  | 
Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>  | 
|
684  | 
\<in> parts (knows Spy evs);  | 
|
685  | 
\<not> expiredK Tk evs;  | 
|
686  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
687  | 
\<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"  | 
|
688  | 
by (blast dest!: Kab_authentic  | 
|
689  | 
intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE])  | 
|
690  | 
||
691  | 
||
692  | 
subsection{*Combined guarantees of key distribution and non-injective agreement on the session keys*}
 | 
|
693  | 
||
694  | 
lemma B_authenticates_and_keydist_to_A:  | 
|
695  | 
"\<lbrakk> Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,  | 
|
696  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs;  | 
|
697  | 
Key K \<notin> analz (spies evs);  | 
|
698  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
699  | 
\<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,  | 
|
700  | 
Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs  | 
|
701  | 
\<and> Key K \<in> analz (knows A evs)"  | 
|
702  | 
apply (blast dest: B_authenticates_A BK3_imp_Gets Gets_A_knows_K)  | 
|
703  | 
done  | 
|
704  | 
||
705  | 
lemma A_authenticates_and_keydist_to_B:  | 
|
706  | 
"\<lbrakk> Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;  | 
|
707  | 
Gets A (Crypt K (Number Ta)) \<in> set evs;  | 
|
708  | 
Key K \<notin> analz (spies evs);  | 
|
709  | 
A \<notin> bad; B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>  | 
|
710  | 
\<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs  | 
|
711  | 
\<and> Key K \<in> analz (knows B evs)"  | 
|
712  | 
apply (blast dest: A_authenticates_B BK4_imp_Gets Gets_B_knows_K)  | 
|
713  | 
done  | 
|
714  | 
||
715  | 
||
716  | 
||
717  | 
||
718  | 
||
719  | 
end  |