| author | wenzelm | 
| Thu, 08 Sep 2022 16:22:44 +0200 | |
| changeset 76086 | 338adf8d423c | 
| parent 69597 | ff784d5a5bfb | 
| child 76287 | cdc14f94c754 | 
| permissions | -rw-r--r-- | 
| 37936 | 1  | 
(* Title: HOL/Auth/OtwayReesBella.thy  | 
| 18886 | 2  | 
Author: Giampaolo Bella, Catania University  | 
3  | 
*)  | 
|
4  | 
||
| 61830 | 5  | 
section\<open>Bella's version of the Otway-Rees protocol\<close>  | 
| 18886 | 6  | 
|
7  | 
||
8  | 
theory OtwayReesBella imports Public begin  | 
|
9  | 
||
| 61830 | 10  | 
text\<open>Bella's modifications to a version of the Otway-Rees protocol taken from  | 
| 18886 | 11  | 
the BAN paper only concern message 7. The updated protocol makes the goal of  | 
12  | 
key distribution of the session key available to A. Investigating the  | 
|
13  | 
principle of Goal Availability undermines the BAN claim about the original  | 
|
14  | 
protocol, that "this protocol does not make use of Kab as an encryption key,  | 
|
15  | 
so neither principal can know whether the key is known to the other". The  | 
|
16  | 
updated protocol makes no use of the session key to encrypt but informs A that  | 
|
| 61830 | 17  | 
B knows it.\<close>  | 
| 18886 | 18  | 
|
| 23746 | 19  | 
inductive_set orb :: "event list set"  | 
20  | 
where  | 
|
| 18886 | 21  | 
|
22  | 
Nil: "[]\<in> orb"  | 
|
23  | 
||
| 23746 | 24  | 
| Fake: "\<lbrakk>evsa\<in> orb; X\<in> synth (analz (knows Spy evsa))\<rbrakk>  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
25  | 
\<Longrightarrow> Says Spy B X # evsa \<in> orb"  | 
| 18886 | 26  | 
|
| 23746 | 27  | 
| Reception: "\<lbrakk>evsr\<in> orb; Says A B X \<in> set evsr\<rbrakk>  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
28  | 
\<Longrightarrow> Gets B X # evsr \<in> orb"  | 
| 18886 | 29  | 
|
| 23746 | 30  | 
| OR1: "\<lbrakk>evs1\<in> orb; Nonce NA \<notin> used evs1\<rbrakk>  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
31  | 
\<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
32  | 
Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
33  | 
# evs1 \<in> orb"  | 
| 18886 | 34  | 
|
| 23746 | 35  | 
| OR2: "\<lbrakk>evs2\<in> orb; Nonce NB \<notin> used evs2;  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
36  | 
Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
37  | 
\<Longrightarrow> Says B Server  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
38  | 
\<lbrace>Nonce M, Agent A, Agent B, X,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
39  | 
Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
40  | 
# evs2 \<in> orb"  | 
| 18886 | 41  | 
|
| 23746 | 42  | 
| OR3: "\<lbrakk>evs3\<in> orb; Key KAB \<notin> used evs3;  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
43  | 
Gets Server  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
44  | 
\<lbrace>Nonce M, Agent A, Agent B,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
45  | 
Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
46  | 
Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
47  | 
\<in> set evs3\<rbrakk>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
48  | 
\<Longrightarrow> Says Server B \<lbrace>Nonce M,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
49  | 
Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
50  | 
Nonce NB, Key KAB\<rbrace>\<rbrace>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
51  | 
# evs3 \<in> orb"  | 
| 18886 | 52  | 
|
53  | 
(*B can only check that the message he is bouncing is a ciphertext*)  | 
|
54  | 
(*Sending M back is omitted*)  | 
|
| 23746 | 55  | 
| OR4: "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>;  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
56  | 
Says B Server \<lbrace>Nonce M, Agent A, Agent B, X',  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
57  | 
Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
58  | 
\<in> set evs4;  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
59  | 
Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
60  | 
\<in> set evs4\<rbrakk>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
61  | 
\<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"  | 
| 18886 | 62  | 
|
63  | 
||
| 23746 | 64  | 
| Oops: "\<lbrakk>evso\<in> orb;  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
65  | 
Says Server B \<lbrace>Nonce M,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
66  | 
Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
67  | 
Nonce NB, Key KAB\<rbrace>\<rbrace>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30549 
diff
changeset
 | 
68  | 
\<in> set evso\<rbrakk>  | 
| 18886 | 69  | 
\<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB, Key KAB\<rbrace> # evso  | 
70  | 
\<in> orb"  | 
|
71  | 
||
72  | 
||
73  | 
||
74  | 
declare knows_Spy_partsEs [elim]  | 
|
75  | 
declare analz_into_parts [dest]  | 
|
76  | 
declare Fake_parts_insert_in_Un [dest]  | 
|
77  | 
||
78  | 
||
| 61830 | 79  | 
text\<open>Fragile proof, with backtracking in the possibility call.\<close>  | 
| 18886 | 80  | 
lemma possibility_thm: "\<lbrakk>A \<noteq> Server; B \<noteq> Server; Key K \<notin> used[]\<rbrakk>  | 
81  | 
\<Longrightarrow> \<exists> evs \<in> orb.  | 
|
82  | 
Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs"  | 
|
83  | 
apply (intro exI bexI)  | 
|
84  | 
apply (rule_tac [2] orb.Nil  | 
|
85  | 
[THEN orb.OR1, THEN orb.Reception,  | 
|
86  | 
THEN orb.OR2, THEN orb.Reception,  | 
|
87  | 
THEN orb.OR3, THEN orb.Reception, THEN orb.OR4])  | 
|
88  | 
apply (possibility, simp add: used_Cons)  | 
|
89  | 
done  | 
|
90  | 
||
91  | 
||
92  | 
lemma Gets_imp_Says :  | 
|
93  | 
"\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"  | 
|
94  | 
apply (erule rev_mp)  | 
|
95  | 
apply (erule orb.induct)  | 
|
96  | 
apply auto  | 
|
97  | 
done  | 
|
98  | 
||
99  | 
lemma Gets_imp_knows_Spy:  | 
|
100  | 
"\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk> \<Longrightarrow> X \<in> knows Spy evs"  | 
|
101  | 
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)  | 
|
102  | 
||
103  | 
declare Gets_imp_knows_Spy [THEN parts.Inj, dest]  | 
|
104  | 
||
105  | 
lemma Gets_imp_knows:  | 
|
106  | 
"\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk> \<Longrightarrow> X \<in> knows B evs"  | 
|
| 
39251
 
8756b44582e2
Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
 
paulson 
parents: 
37936 
diff
changeset
 | 
107  | 
by (metis Gets_imp_knows_Spy Gets_imp_knows_agents)  | 
| 18886 | 108  | 
|
109  | 
lemma OR2_analz_knows_Spy:  | 
|
110  | 
"\<lbrakk>Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs; evs \<in> orb\<rbrakk>  | 
|
111  | 
\<Longrightarrow> X \<in> analz (knows Spy evs)"  | 
|
112  | 
by (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])  | 
|
113  | 
||
114  | 
lemma OR4_parts_knows_Spy:  | 
|
115  | 
"\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key Kab\<rbrace>\<rbrace> \<in> set evs;  | 
|
116  | 
evs \<in> orb\<rbrakk> \<Longrightarrow> X \<in> parts (knows Spy evs)"  | 
|
117  | 
by blast  | 
|
118  | 
||
119  | 
lemma Oops_parts_knows_Spy:  | 
|
120  | 
"Says Server B \<lbrace>Nonce M, Crypt K' \<lbrace>X, Nonce Nb, K\<rbrace>\<rbrace> \<in> set evs  | 
|
121  | 
\<Longrightarrow> K \<in> parts (knows Spy evs)"  | 
|
122  | 
by blast  | 
|
123  | 
||
124  | 
lemmas OR2_parts_knows_Spy =  | 
|
| 45605 | 125  | 
OR2_analz_knows_Spy [THEN analz_into_parts]  | 
| 18886 | 126  | 
|
127  | 
ML  | 
|
| 61830 | 128  | 
\<open>  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
129  | 
fun parts_explicit_tac ctxt i =  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
130  | 
    forward_tac ctxt [@{thm Oops_parts_knows_Spy}] (i+7) THEN
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
131  | 
    forward_tac ctxt [@{thm OR4_parts_knows_Spy}]  (i+6) THEN
 | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
132  | 
    forward_tac ctxt [@{thm OR2_parts_knows_Spy}]  (i+4)
 | 
| 61830 | 133  | 
\<close>  | 
| 18886 | 134  | 
|
| 61830 | 135  | 
method_setup parts_explicit = \<open>  | 
136  | 
Scan.succeed (SIMPLE_METHOD' o parts_explicit_tac)\<close>  | 
|
| 18886 | 137  | 
"to explicitly state that some message components belong to parts knows Spy"  | 
138  | 
||
139  | 
||
140  | 
lemma Spy_see_shrK [simp]:  | 
|
141  | 
"evs \<in> orb \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"  | 
|
142  | 
by (erule orb.induct, parts_explicit, simp_all, blast+)  | 
|
143  | 
||
144  | 
lemma Spy_analz_shrK [simp]:  | 
|
145  | 
"evs \<in> orb \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"  | 
|
146  | 
by auto  | 
|
147  | 
||
148  | 
lemma Spy_see_shrK_D [dest!]:  | 
|
149  | 
"[|Key (shrK A) \<in> parts (knows Spy evs); evs \<in> orb|] ==> A \<in> bad"  | 
|
150  | 
by (blast dest: Spy_see_shrK)  | 
|
151  | 
||
152  | 
lemma new_keys_not_used [simp]:  | 
|
153  | 
"\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> orb\<rbrakk> \<Longrightarrow> K \<notin> keysFor (parts (knows Spy evs))"  | 
|
154  | 
apply (erule rev_mp)  | 
|
155  | 
apply (erule orb.induct, parts_explicit, simp_all)  | 
|
156  | 
apply (force dest!: keysFor_parts_insert)  | 
|
157  | 
apply (blast+)  | 
|
158  | 
done  | 
|
159  | 
||
160  | 
||
161  | 
||
| 61830 | 162  | 
subsection\<open>Proofs involving analz\<close>  | 
| 18886 | 163  | 
|
| 61830 | 164  | 
text\<open>Describes the form of K and NA when the Server sends this message. Also  | 
165  | 
for Oops case.\<close>  | 
|
| 18886 | 166  | 
lemma Says_Server_message_form:  | 
167  | 
"\<lbrakk>Says Server B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  | 
|
168  | 
evs \<in> orb\<rbrakk>  | 
|
| 67613 | 169  | 
\<Longrightarrow> K \<notin> range shrK \<and> (\<exists> A Na. X=(Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>))"  | 
| 18886 | 170  | 
by (erule rev_mp, erule orb.induct, simp_all)  | 
171  | 
||
172  | 
lemma Says_Server_imp_Gets:  | 
|
173  | 
"\<lbrakk>Says Server B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>,  | 
|
174  | 
Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  | 
|
175  | 
evs \<in> orb\<rbrakk>  | 
|
176  | 
\<Longrightarrow> Gets Server \<lbrace>Nonce M, Agent A, Agent B,  | 
|
177  | 
Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>,  | 
|
178  | 
Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>  | 
|
179  | 
\<in> set evs"  | 
|
180  | 
by (erule rev_mp, erule orb.induct, simp_all)  | 
|
181  | 
||
182  | 
||
183  | 
lemma A_trusts_OR1:  | 
|
184  | 
"\<lbrakk>Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs);  | 
|
185  | 
A \<notin> bad; evs \<in> orb\<rbrakk>  | 
|
186  | 
\<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs"  | 
|
187  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)  | 
|
188  | 
apply (blast)  | 
|
189  | 
done  | 
|
190  | 
||
191  | 
||
192  | 
lemma B_trusts_OR2:  | 
|
193  | 
"\<lbrakk>Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>  | 
|
194  | 
\<in> parts (knows Spy evs); B \<notin> bad; evs \<in> orb\<rbrakk>  | 
|
195  | 
\<Longrightarrow> (\<exists> X. Says B Server \<lbrace>Nonce M, Agent A, Agent B, X,  | 
|
196  | 
Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>  | 
|
197  | 
\<in> set evs)"  | 
|
198  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)  | 
|
199  | 
apply (blast+)  | 
|
200  | 
done  | 
|
201  | 
||
202  | 
||
203  | 
lemma B_trusts_OR3:  | 
|
204  | 
"\<lbrakk>Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace> \<in> parts (knows Spy evs);  | 
|
205  | 
B \<notin> bad; evs \<in> orb\<rbrakk>  | 
|
206  | 
\<Longrightarrow> \<exists> M. Says Server B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace>  | 
|
207  | 
\<in> set evs"  | 
|
208  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)  | 
|
209  | 
apply (blast+)  | 
|
210  | 
done  | 
|
211  | 
||
212  | 
lemma Gets_Server_message_form:  | 
|
213  | 
"\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  | 
|
214  | 
evs \<in> orb\<rbrakk>  | 
|
| 67613 | 215  | 
\<Longrightarrow> (K \<notin> range shrK \<and> (\<exists> A Na. X = (Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>)))  | 
| 18886 | 216  | 
| X \<in> analz (knows Spy evs)"  | 
| 
39251
 
8756b44582e2
Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
 
paulson 
parents: 
37936 
diff
changeset
 | 
217  | 
by (metis B_trusts_OR3 Crypt_Spy_analz_bad Gets_imp_Says MPair_analz MPair_parts  | 
| 
 
8756b44582e2
Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
 
paulson 
parents: 
37936 
diff
changeset
 | 
218  | 
Says_Server_message_form Says_imp_analz_Spy Says_imp_parts_knows_Spy)  | 
| 18886 | 219  | 
|
220  | 
lemma unique_Na: "\<lbrakk>Says A B \<lbrace>Nonce M, Agent A, Agent B, Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs;  | 
|
221  | 
Says A B' \<lbrace>Nonce M', Agent A, Agent B', Crypt (shrK A) \<lbrace>Nonce Na, Nonce M', Agent A, Agent B'\<rbrace>\<rbrace> \<in> set evs;  | 
|
| 67613 | 222  | 
A \<notin> bad; evs \<in> orb\<rbrakk> \<Longrightarrow> B=B' \<and> M=M'"  | 
| 18886 | 223  | 
by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+)  | 
224  | 
||
225  | 
lemma unique_Nb: "\<lbrakk>Says B Server \<lbrace>Nonce M, Agent A, Agent B, X, Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs;  | 
|
226  | 
Says B Server \<lbrace>Nonce M', Agent A', Agent B, X', Crypt (shrK B) \<lbrace>Nonce Nb,Nonce M', Nonce M', Agent A', Agent B\<rbrace>\<rbrace> \<in> set evs;  | 
|
| 67613 | 227  | 
B \<notin> bad; evs \<in> orb\<rbrakk> \<Longrightarrow> M=M' \<and> A=A' \<and> X=X'"  | 
| 18886 | 228  | 
by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+)  | 
229  | 
||
230  | 
lemma analz_image_freshCryptK_lemma:  | 
|
231  | 
"(Crypt K X \<in> analz (Key`nE \<union> H)) \<longrightarrow> (Crypt K X \<in> analz H) \<Longrightarrow>  | 
|
| 
39251
 
8756b44582e2
Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
 
paulson 
parents: 
37936 
diff
changeset
 | 
232  | 
(Crypt K X \<in> analz (Key`nE \<union> H)) = (Crypt K X \<in> analz H)"  | 
| 18886 | 233  | 
by (blast intro: analz_mono [THEN [2] rev_subsetD])  | 
234  | 
||
235  | 
ML  | 
|
| 61830 | 236  | 
\<open>  | 
| 24122 | 237  | 
structure OtwayReesBella =  | 
238  | 
struct  | 
|
| 18886 | 239  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45605 
diff
changeset
 | 
240  | 
val analz_image_freshK_ss =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45605 
diff
changeset
 | 
241  | 
simpset_of  | 
| 69597 | 242  | 
(\<^context> delsimps [image_insert, image_Un]  | 
| 24122 | 243  | 
      delsimps [@{thm imp_disjL}]    (*reduces blow-up*)
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45605 
diff
changeset
 | 
244  | 
      addsimps @{thms analz_image_freshK_simps})
 | 
| 24122 | 245  | 
|
246  | 
end  | 
|
| 61830 | 247  | 
\<close>  | 
| 18886 | 248  | 
|
| 61830 | 249  | 
method_setup analz_freshCryptK = \<open>  | 
| 30549 | 250  | 
Scan.succeed (fn ctxt =>  | 
| 
30510
 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 
wenzelm 
parents: 
24122 
diff
changeset
 | 
251  | 
(SIMPLE_METHOD  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
252  | 
(EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),  | 
| 60754 | 253  | 
          REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshCryptK_lemma}),
 | 
| 24122 | 254  | 
ALLGOALS (asm_simp_tac  | 
| 61830 | 255  | 
(put_simpset OtwayReesBella.analz_image_freshK_ss ctxt))])))\<close>  | 
| 18886 | 256  | 
"for proving useful rewrite rule"  | 
257  | 
||
258  | 
||
| 61830 | 259  | 
method_setup disentangle = \<open>  | 
| 30549 | 260  | 
Scan.succeed  | 
| 51798 | 261  | 
(fn ctxt => SIMPLE_METHOD  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
262  | 
(REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE, disjE]  | 
| 61830 | 263  | 
ORELSE' hyp_subst_tac ctxt)))\<close>  | 
| 18886 | 264  | 
"for eliminating conjunctions, disjunctions and the like"  | 
265  | 
||
266  | 
||
267  | 
||
268  | 
lemma analz_image_freshCryptK [rule_format]:  | 
|
269  | 
"evs \<in> orb \<Longrightarrow>  | 
|
270  | 
Key K \<notin> analz (knows Spy evs) \<longrightarrow>  | 
|
271  | 
(\<forall> KK. KK \<subseteq> - (range shrK) \<longrightarrow>  | 
|
272  | 
(Crypt K X \<in> analz (Key`KK \<union> (knows Spy evs))) =  | 
|
273  | 
(Crypt K X \<in> analz (knows Spy evs)))"  | 
|
274  | 
apply (erule orb.induct)  | 
|
275  | 
apply (analz_mono_contra)  | 
|
276  | 
apply (frule_tac [7] Gets_Server_message_form)  | 
|
277  | 
apply (frule_tac [9] Says_Server_message_form)  | 
|
278  | 
apply disentangle  | 
|
279  | 
apply (drule_tac [5] Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Snd])  | 
|
280  | 
prefer 8 apply clarify  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
39251 
diff
changeset
 | 
281  | 
apply (analz_freshCryptK, spy_analz, fastforce)  | 
| 18886 | 282  | 
done  | 
283  | 
||
284  | 
||
285  | 
||
286  | 
lemma analz_insert_freshCryptK:  | 
|
287  | 
"\<lbrakk>evs \<in> orb; Key K \<notin> analz (knows Spy evs);  | 
|
288  | 
Seskey \<notin> range shrK\<rbrakk> \<Longrightarrow>  | 
|
289  | 
(Crypt K X \<in> analz (insert (Key Seskey) (knows Spy evs))) =  | 
|
290  | 
(Crypt K X \<in> analz (knows Spy evs))"  | 
|
291  | 
by (simp only: analz_image_freshCryptK analz_image_freshK_simps)  | 
|
292  | 
||
293  | 
||
294  | 
lemma analz_hard:  | 
|
295  | 
"\<lbrakk>Says A B \<lbrace>Nonce M, Agent A, Agent B,  | 
|
296  | 
Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in>set evs;  | 
|
297  | 
Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace> \<in> analz (knows Spy evs);  | 
|
298  | 
A \<notin> bad; B \<notin> bad; evs \<in> orb\<rbrakk>  | 
|
299  | 
\<Longrightarrow> Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs"  | 
|
300  | 
apply (erule rev_mp)  | 
|
301  | 
apply (erule rev_mp)  | 
|
302  | 
apply (erule orb.induct)  | 
|
303  | 
apply (frule_tac [7] Gets_Server_message_form)  | 
|
304  | 
apply (frule_tac [9] Says_Server_message_form)  | 
|
305  | 
apply disentangle  | 
|
| 61830 | 306  | 
txt\<open>letting the simplifier solve OR2\<close>  | 
| 18886 | 307  | 
apply (drule_tac [5] Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Snd])  | 
308  | 
apply (simp_all (no_asm_simp) add: analz_insert_eq pushes split_ifs)  | 
|
309  | 
apply (spy_analz)  | 
|
| 61830 | 310  | 
txt\<open>OR1\<close>  | 
| 18886 | 311  | 
apply blast  | 
| 61830 | 312  | 
txt\<open>Oops\<close>  | 
| 18886 | 313  | 
prefer 4 apply (blast dest: analz_insert_freshCryptK)  | 
| 61830 | 314  | 
txt\<open>OR4 - ii\<close>  | 
| 18886 | 315  | 
prefer 3 apply (blast)  | 
| 61830 | 316  | 
txt\<open>OR3\<close>  | 
| 18886 | 317  | 
(*adding Gets_imp_ and Says_imp_ for efficiency*)  | 
318  | 
apply (blast dest:  | 
|
319  | 
A_trusts_OR1 unique_Na Key_not_used analz_insert_freshCryptK)  | 
|
| 61830 | 320  | 
txt\<open>OR4 - i\<close>  | 
| 18886 | 321  | 
apply clarify  | 
322  | 
apply (simp add: pushes split_ifs)  | 
|
323  | 
apply (case_tac "Aaa\<in>bad")  | 
|
324  | 
apply (blast dest: analz_insert_freshCryptK)  | 
|
325  | 
apply clarify  | 
|
326  | 
apply simp  | 
|
327  | 
apply (case_tac "Ba\<in>bad")  | 
|
328  | 
apply (frule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Decrypt, THEN analz.Fst] , assumption)  | 
|
329  | 
apply (simp (no_asm_simp))  | 
|
330  | 
apply clarify  | 
|
331  | 
apply (frule Gets_imp_knows_Spy  | 
|
332  | 
[THEN parts.Inj, THEN parts.Snd, THEN B_trusts_OR3],  | 
|
333  | 
assumption, assumption, assumption, erule exE)  | 
|
334  | 
apply (frule Says_Server_imp_Gets  | 
|
335  | 
[THEN Gets_imp_knows_Spy, THEN parts.Inj, THEN parts.Snd,  | 
|
336  | 
THEN parts.Snd, THEN parts.Snd, THEN parts.Fst, THEN A_trusts_OR1],  | 
|
337  | 
assumption, assumption, assumption, assumption)  | 
|
338  | 
apply (blast dest: Says_Server_imp_Gets B_trusts_OR2 unique_Na unique_Nb)  | 
|
339  | 
done  | 
|
340  | 
||
341  | 
||
342  | 
lemma Gets_Server_message_form':  | 
|
343  | 
"\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  | 
|
344  | 
B \<notin> bad; evs \<in> orb\<rbrakk>  | 
|
| 67613 | 345  | 
\<Longrightarrow> K \<notin> range shrK \<and> (\<exists> A Na. X = (Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>))"  | 
| 18886 | 346  | 
by (blast dest!: B_trusts_OR3 Says_Server_message_form)  | 
347  | 
||
348  | 
||
349  | 
lemma OR4_imp_Gets:  | 
|
350  | 
"\<lbrakk>Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs;  | 
|
351  | 
B \<notin> bad; evs \<in> orb\<rbrakk>  | 
|
352  | 
\<Longrightarrow> (\<exists> Nb. Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>,  | 
|
353  | 
Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs)"  | 
|
354  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)  | 
|
355  | 
prefer 3 apply (blast dest: Gets_Server_message_form')  | 
|
356  | 
apply blast+  | 
|
357  | 
done  | 
|
358  | 
||
359  | 
||
360  | 
lemma A_keydist_to_B:  | 
|
361  | 
"\<lbrakk>Says A B \<lbrace>Nonce M, Agent A, Agent B,  | 
|
362  | 
Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in>set evs;  | 
|
363  | 
Gets A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs;  | 
|
364  | 
A \<notin> bad; B \<notin> bad; evs \<in> orb\<rbrakk>  | 
|
365  | 
\<Longrightarrow> Key K \<in> analz (knows B evs)"  | 
|
366  | 
apply (drule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd], assumption)  | 
|
367  | 
apply (drule analz_hard, assumption, assumption, assumption, assumption)  | 
|
368  | 
apply (drule OR4_imp_Gets, assumption, assumption)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
39251 
diff
changeset
 | 
369  | 
apply (fastforce dest!: Gets_imp_knows [THEN analz.Inj] analz.Decrypt)  | 
| 18886 | 370  | 
done  | 
371  | 
||
372  | 
||
| 61830 | 373  | 
text\<open>Other properties as for the original protocol\<close>  | 
| 18886 | 374  | 
|
375  | 
||
376  | 
end  |