| author | mueller | 
| Tue, 20 May 1997 10:28:19 +0200 | |
| changeset 3225 | cee363fc07d7 | 
| parent 3207 | fe79ad367d77 | 
| child 3451 | d10f100676d8 | 
| permissions | -rw-r--r-- | 
| 2090 | 1  | 
(* Title: HOL/Auth/OtwayRees  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1996 University of Cambridge  | 
|
5  | 
||
6  | 
Inductive relation "otway" for the Otway-Rees protocol.  | 
|
7  | 
||
8  | 
Simplified version with minimal encryption but explicit messages  | 
|
9  | 
||
10  | 
From page 11 of  | 
|
11  | 
Abadi and Needham. Prudent Engineering Practice for Cryptographic Protocols.  | 
|
12  | 
IEEE Trans. SE 22 (1), 1996  | 
|
13  | 
*)  | 
|
14  | 
||
15  | 
open OtwayRees_AN;  | 
|
16  | 
||
17  | 
proof_timing:=true;  | 
|
18  | 
HOL_quantifiers := false;  | 
|
19  | 
||
20  | 
||
| 2331 | 21  | 
(*A "possibility property": there are traces that reach the end*)  | 
| 2090 | 22  | 
goal thy  | 
| 2331 | 23  | 
"!!A B. [| A ~= B; A ~= Server; B ~= Server |] \  | 
24  | 
\ ==> EX K. EX NA. EX evs: otway lost. \  | 
|
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
25  | 
\             Says B A (Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|}) \
 | 
| 2090 | 26  | 
\ : set_of_list evs";  | 
27  | 
by (REPEAT (resolve_tac [exI,bexI] 1));  | 
|
28  | 
by (rtac (otway.Nil RS otway.OR1 RS otway.OR2 RS otway.OR3 RS otway.OR4) 2);  | 
|
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
29  | 
by possibility_tac;  | 
| 2090 | 30  | 
result();  | 
31  | 
||
32  | 
||
33  | 
(**** Inductive proofs about otway ****)  | 
|
34  | 
||
| 2106 | 35  | 
(*Monotonicity*)  | 
| 2090 | 36  | 
goal thy "!!evs. lost' <= lost ==> otway lost' <= otway lost";  | 
37  | 
by (rtac subsetI 1);  | 
|
38  | 
by (etac otway.induct 1);  | 
|
39  | 
by (REPEAT_FIRST  | 
|
40  | 
(best_tac (!claset addIs (impOfSubs (sees_mono RS analz_mono RS synth_mono)  | 
|
41  | 
:: otway.intrs))));  | 
|
42  | 
qed "otway_mono";  | 
|
43  | 
||
44  | 
(*Nobody sends themselves messages*)  | 
|
45  | 
goal thy "!!evs. evs : otway lost ==> ALL A X. Says A A X ~: set_of_list evs";  | 
|
46  | 
by (etac otway.induct 1);  | 
|
47  | 
by (Auto_tac());  | 
|
48  | 
qed_spec_mp "not_Says_to_self";  | 
|
49  | 
Addsimps [not_Says_to_self];  | 
|
50  | 
AddSEs [not_Says_to_self RSN (2, rev_notE)];  | 
|
51  | 
||
52  | 
||
53  | 
(** For reasoning about the encrypted portion of messages **)  | 
|
54  | 
||
| 
2837
 
dee1c7f1f576
Trivial renamings (for consistency with CSFW papers)
 
paulson 
parents: 
2637 
diff
changeset
 | 
55  | 
goal thy "!!evs. Says S' B {|X, Crypt(shrK B) X'|} : set_of_list evs ==> \
 | 
| 2090 | 56  | 
\ X : analz (sees lost Spy evs)";  | 
| 3102 | 57  | 
by (blast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]) 1);  | 
| 2090 | 58  | 
qed "OR4_analz_sees_Spy";  | 
59  | 
||
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
60  | 
goal thy "!!evs. Says Server B {|X, Crypt K' {|NB, a, Agent B, K|}|} \
 | 
| 2131 | 61  | 
\ : set_of_list evs ==> K : parts (sees lost Spy evs)";  | 
| 3102 | 62  | 
by (blast_tac (!claset addSEs sees_Spy_partsEs) 1);  | 
| 2131 | 63  | 
qed "Oops_parts_sees_Spy";  | 
| 2090 | 64  | 
|
| 2131 | 65  | 
(*OR4_analz_sees_Spy lets us treat those cases using the same  | 
| 2090 | 66  | 
argument as for the Fake case. This is possible for most, but not all,  | 
| 2131 | 67  | 
proofs, since Fake messages originate from the Spy. *)  | 
| 2090 | 68  | 
|
69  | 
bind_thm ("OR4_parts_sees_Spy",
 | 
|
70  | 
OR4_analz_sees_Spy RS (impOfSubs analz_subset_parts));  | 
|
71  | 
||
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
72  | 
(*For proving the easier theorems about X ~: parts (sees lost Spy evs).  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
73  | 
We instantiate the variable to "lost" since leaving it as a Var would  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
74  | 
interfere with simplification.*)  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
75  | 
val parts_induct_tac =  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
76  | 
    let val tac = forw_inst_tac [("lost","lost")] 
 | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
77  | 
in etac otway.induct 1 THEN  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
78  | 
tac OR4_parts_sees_Spy 6 THEN  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
79  | 
tac Oops_parts_sees_Spy 7 THEN  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
80  | 
prove_simple_subgoals_tac 1  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
81  | 
end;  | 
| 2090 | 82  | 
|
83  | 
||
84  | 
(** Theorems of the form X ~: parts (sees lost Spy evs) imply that NOBODY  | 
|
85  | 
sends messages containing X! **)  | 
|
86  | 
||
87  | 
(*Spy never sees another agent's shared key! (unless it's lost at start)*)  | 
|
88  | 
goal thy  | 
|
| 2131 | 89  | 
"!!evs. evs : otway lost \  | 
90  | 
\ ==> (Key (shrK A) : parts (sees lost Spy evs)) = (A : lost)";  | 
|
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
91  | 
by parts_induct_tac;  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
92  | 
by (Fake_parts_insert_tac 1);  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
93  | 
by (Blast_tac 1);  | 
| 2131 | 94  | 
qed "Spy_see_shrK";  | 
95  | 
Addsimps [Spy_see_shrK];  | 
|
| 2090 | 96  | 
|
| 2131 | 97  | 
goal thy  | 
98  | 
"!!evs. evs : otway lost \  | 
|
99  | 
\ ==> (Key (shrK A) : analz (sees lost Spy evs)) = (A : lost)";  | 
|
100  | 
by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset));  | 
|
101  | 
qed "Spy_analz_shrK";  | 
|
102  | 
Addsimps [Spy_analz_shrK];  | 
|
| 2090 | 103  | 
|
| 2131 | 104  | 
goal thy "!!A. [| Key (shrK A) : parts (sees lost Spy evs); \  | 
105  | 
\ evs : otway lost |] ==> A:lost";  | 
|
| 3102 | 106  | 
by (blast_tac (!claset addDs [Spy_see_shrK]) 1);  | 
| 2131 | 107  | 
qed "Spy_see_shrK_D";  | 
| 2090 | 108  | 
|
| 2131 | 109  | 
bind_thm ("Spy_analz_shrK_D", analz_subset_parts RS subsetD RS Spy_see_shrK_D);
 | 
110  | 
AddSDs [Spy_see_shrK_D, Spy_analz_shrK_D];  | 
|
| 2090 | 111  | 
|
112  | 
||
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
113  | 
(*Nobody can have used non-existent keys!*)  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
114  | 
goal thy "!!evs. evs : otway lost ==> \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
115  | 
\ Key K ~: used evs --> K ~: keysFor (parts (sees lost Spy evs))";  | 
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
116  | 
by parts_induct_tac;  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
117  | 
(*Fake*)  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
118  | 
by (best_tac  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
119  | 
(!claset addIs [impOfSubs analz_subset_parts]  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
120  | 
addDs [impOfSubs (analz_subset_parts RS keysFor_mono),  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
121  | 
impOfSubs (parts_insert_subset_Un RS keysFor_mono)]  | 
| 3207 | 122  | 
addss (!simpset)) 1);  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
123  | 
(*OR3*)  | 
| 3102 | 124  | 
by (Blast_tac 1);  | 
| 2160 | 125  | 
qed_spec_mp "new_keys_not_used";  | 
| 2090 | 126  | 
|
127  | 
bind_thm ("new_keys_not_analzd",
 | 
|
128  | 
[analz_subset_parts RS keysFor_mono,  | 
|
129  | 
new_keys_not_used] MRS contra_subsetD);  | 
|
130  | 
||
131  | 
Addsimps [new_keys_not_used, new_keys_not_analzd];  | 
|
132  | 
||
133  | 
||
134  | 
||
135  | 
(*** Proofs involving analz ***)  | 
|
136  | 
||
| 2131 | 137  | 
(*Describes the form of K and NA when the Server sends this message.*)  | 
| 2090 | 138  | 
goal thy  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
139  | 
"!!evs. [| Says Server B \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
140  | 
\              {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},    \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
141  | 
\                Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}   \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
142  | 
\ : set_of_list evs; \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
143  | 
\ evs : otway lost |] \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
144  | 
\ ==> K ~: range shrK & (EX i. NA = Nonce i) & (EX j. NB = Nonce j)";  | 
| 2131 | 145  | 
by (etac rev_mp 1);  | 
146  | 
by (etac otway.induct 1);  | 
|
| 3102 | 147  | 
by (ALLGOALS Asm_simp_tac);  | 
148  | 
by (Blast_tac 1);  | 
|
| 2131 | 149  | 
qed "Says_Server_message_form";  | 
| 2090 | 150  | 
|
151  | 
||
152  | 
(*For proofs involving analz. We again instantiate the variable to "lost".*)  | 
|
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
153  | 
val analz_sees_tac =  | 
| 2090 | 154  | 
    dres_inst_tac [("lost","lost")] OR4_analz_sees_Spy 6 THEN
 | 
| 2131 | 155  | 
    forw_inst_tac [("lost","lost")] Says_Server_message_form 7 THEN
 | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
156  | 
assume_tac 7 THEN  | 
| 
2451
 
ce85a2aafc7a
Extensive tidying and simplification, largely stemming from
 
paulson 
parents: 
2417 
diff
changeset
 | 
157  | 
REPEAT ((eresolve_tac [exE, conjE] ORELSE' hyp_subst_tac) 7);  | 
| 2090 | 158  | 
|
159  | 
||
160  | 
(****  | 
|
161  | 
The following is to prove theorems of the form  | 
|
162  | 
||
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
163  | 
Key K : analz (insert (Key KAB) (sees lost Spy evs)) ==>  | 
| 
2451
 
ce85a2aafc7a
Extensive tidying and simplification, largely stemming from
 
paulson 
parents: 
2417 
diff
changeset
 | 
164  | 
Key K : analz (sees lost Spy evs)  | 
| 2090 | 165  | 
|
166  | 
A more general formula must be proved inductively.  | 
|
167  | 
****)  | 
|
168  | 
||
169  | 
||
170  | 
(** Session keys are not used to encrypt other session keys **)  | 
|
171  | 
||
172  | 
(*The equality makes the induction hypothesis easier to apply*)  | 
|
173  | 
goal thy  | 
|
| 2331 | 174  | 
"!!evs. evs : otway lost ==> \  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
175  | 
\ ALL K KK. KK <= Compl (range shrK) --> \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
176  | 
\ (Key K : analz (Key``KK Un (sees lost Spy evs))) = \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
177  | 
\ (K : KK | Key K : analz (sees lost Spy evs))";  | 
| 2090 | 178  | 
by (etac otway.induct 1);  | 
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
179  | 
by analz_sees_tac;  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
180  | 
by (REPEAT_FIRST (resolve_tac [allI, impI]));  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
181  | 
by (REPEAT_FIRST (rtac analz_image_freshK_lemma ));  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
182  | 
(*14 seconds*)  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
183  | 
by (ALLGOALS (asm_simp_tac analz_image_freshK_ss));  | 
| 2131 | 184  | 
(*OR4, Fake*)  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
185  | 
by (EVERY (map spy_analz_tac [3,2]));  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
186  | 
(*Base*)  | 
| 3102 | 187  | 
by (Blast_tac 1);  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
188  | 
qed_spec_mp "analz_image_freshK";  | 
| 2090 | 189  | 
|
190  | 
||
191  | 
goal thy  | 
|
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
192  | 
"!!evs. [| evs : otway lost; KAB ~: range shrK |] ==> \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
193  | 
\ Key K : analz (insert (Key KAB) (sees lost Spy evs)) = \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
194  | 
\ (K = KAB | Key K : analz (sees lost Spy evs))";  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
195  | 
by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
196  | 
qed "analz_insert_freshK";  | 
| 2090 | 197  | 
|
198  | 
||
199  | 
(*** The Key K uniquely identifies the Server's message. **)  | 
|
200  | 
||
201  | 
goal thy  | 
|
| 2331 | 202  | 
"!!evs. evs : otway lost ==> \  | 
| 2090 | 203  | 
\ EX A' B' NA' NB'. ALL A B NA NB. \  | 
204  | 
\ Says Server B \  | 
|
| 2331 | 205  | 
\         {|Crypt (shrK A) {|NA, Agent A, Agent B, K|},                     \
 | 
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
206  | 
\           Crypt (shrK B) {|NB, Agent A, Agent B, K|}|} : set_of_list evs  \
 | 
| 2090 | 207  | 
\ --> A=A' & B=B' & NA=NA' & NB=NB'";  | 
208  | 
by (etac otway.induct 1);  | 
|
209  | 
by (ALLGOALS (asm_simp_tac (!simpset addsimps [all_conj_distrib])));  | 
|
210  | 
by (Step_tac 1);  | 
|
211  | 
(*Remaining cases: OR3 and OR4*)  | 
|
212  | 
by (ex_strip_tac 2);  | 
|
| 3102 | 213  | 
by (Blast_tac 2);  | 
| 2090 | 214  | 
by (expand_case_tac "K = ?y" 1);  | 
215  | 
by (REPEAT (ares_tac [refl,exI,impI,conjI] 2));  | 
|
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
216  | 
(*...we assume X is a recent message and handle this case by contradiction*)  | 
| 3102 | 217  | 
by (blast_tac (!claset addSEs sees_Spy_partsEs  | 
218  | 
delrules[conjI] (*prevent splitup into 4 subgoals*)) 1);  | 
|
| 2090 | 219  | 
val lemma = result();  | 
220  | 
||
221  | 
||
222  | 
goal thy  | 
|
223  | 
"!!evs. [| Says Server B \  | 
|
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
224  | 
\            {|Crypt (shrK A) {|NA, Agent A, Agent B, K|},         \
 | 
| 
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
225  | 
\              Crypt (shrK B) {|NB, Agent A, Agent B, K|}|}        \
 | 
| 2090 | 226  | 
\ : set_of_list evs; \  | 
227  | 
\ Says Server B' \  | 
|
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
228  | 
\            {|Crypt (shrK A') {|NA', Agent A', Agent B', K|},     \
 | 
| 
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
229  | 
\              Crypt (shrK B') {|NB', Agent A', Agent B', K|}|}    \
 | 
| 2090 | 230  | 
\ : set_of_list evs; \  | 
231  | 
\ evs : otway lost |] \  | 
|
232  | 
\ ==> A=A' & B=B' & NA=NA' & NB=NB'";  | 
|
| 2417 | 233  | 
by (prove_unique_tac lemma 1);  | 
| 2090 | 234  | 
qed "unique_session_keys";  | 
235  | 
||
236  | 
||
237  | 
||
238  | 
(**** Authenticity properties relating to NA ****)  | 
|
239  | 
||
240  | 
(*If the encrypted message appears then it originated with the Server!*)  | 
|
241  | 
goal thy  | 
|
242  | 
"!!evs. [| A ~: lost; evs : otway lost |] \  | 
|
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
243  | 
\ ==> Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}        \
 | 
| 2090 | 244  | 
\ : parts (sees lost Spy evs) \  | 
| 2331 | 245  | 
\ --> (EX NB. Says Server B \  | 
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
246  | 
\                  {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},     \
 | 
| 
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
247  | 
\                    Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}    \
 | 
| 2090 | 248  | 
\ : set_of_list evs)";  | 
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
249  | 
by parts_induct_tac;  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
250  | 
by (Fake_parts_insert_tac 1);  | 
| 2090 | 251  | 
by (ALLGOALS (asm_simp_tac (!simpset addsimps [ex_disj_distrib])));  | 
252  | 
(*OR3*)  | 
|
| 3102 | 253  | 
by (Blast_tac 1);  | 
| 2090 | 254  | 
qed_spec_mp "NA_Crypt_imp_Server_msg";  | 
255  | 
||
256  | 
||
| 2454 | 257  | 
(*Corollary: if A receives B's OR4 message then it originated with the Server.  | 
258  | 
Freshness may be inferred from nonce NA.*)  | 
|
| 2090 | 259  | 
goal thy  | 
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
260  | 
 "!!evs. [| Says B' A (Crypt (shrK A) {|NA, Agent A, Agent B, Key K|})  \
 | 
| 2090 | 261  | 
\ : set_of_list evs; \  | 
262  | 
\ A ~: lost; evs : otway lost |] \  | 
|
263  | 
\ ==> EX NB. Says Server B \  | 
|
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
264  | 
\                    {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},  \
 | 
| 
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
265  | 
\                      Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} \
 | 
| 2090 | 266  | 
\ : set_of_list evs";  | 
| 3102 | 267  | 
by (blast_tac (!claset addSIs [NA_Crypt_imp_Server_msg]  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
268  | 
addEs sees_Spy_partsEs) 1);  | 
| 2331 | 269  | 
qed "A_trusts_OR4";  | 
| 2090 | 270  | 
|
271  | 
||
272  | 
(** Crucial secrecy property: Spy does not see the keys sent in msg OR3  | 
|
273  | 
Does not in itself guarantee security: an attack could violate  | 
|
274  | 
the premises, e.g. by having A=Spy **)  | 
|
275  | 
||
276  | 
goal thy  | 
|
| 2166 | 277  | 
"!!evs. [| A ~: lost; B ~: lost; evs : otway lost |] \  | 
| 2090 | 278  | 
\ ==> Says Server B \  | 
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
279  | 
\             {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},            \
 | 
| 
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
280  | 
\               Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}           \
 | 
| 2090 | 281  | 
\ : set_of_list evs --> \  | 
| 2131 | 282  | 
\            Says B Spy {|NA, NB, Key K|} ~: set_of_list evs -->           \
 | 
| 2090 | 283  | 
\ Key K ~: analz (sees lost Spy evs)";  | 
284  | 
by (etac otway.induct 1);  | 
|
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
285  | 
by analz_sees_tac;  | 
| 2090 | 286  | 
by (ALLGOALS  | 
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
287  | 
(asm_simp_tac (!simpset addcongs [conj_cong]  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
288  | 
addsimps [not_parts_not_analz,  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
289  | 
analz_insert_freshK]  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
290  | 
setloop split_tac [expand_if])));  | 
| 2090 | 291  | 
(*OR3*)  | 
| 3102 | 292  | 
by (blast_tac (!claset addSEs sees_Spy_partsEs  | 
293  | 
addIs [impOfSubs analz_subset_parts]) 2);  | 
|
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
294  | 
(*Oops*)  | 
| 3102 | 295  | 
by (blast_tac (!claset addDs [unique_session_keys]) 3);  | 
| 2131 | 296  | 
(*OR4, Fake*)  | 
| 2375 | 297  | 
by (REPEAT_FIRST spy_analz_tac);  | 
| 2090 | 298  | 
val lemma = result() RS mp RS mp RSN(2,rev_notE);  | 
299  | 
||
300  | 
goal thy  | 
|
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
301  | 
"!!evs. [| Says Server B \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
302  | 
\              {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},    \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
303  | 
\                Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}   \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
304  | 
\ : set_of_list evs; \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
305  | 
\           Says B Spy {|NA, NB, Key K|} ~: set_of_list evs;        \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
306  | 
\ A ~: lost; B ~: lost; evs : otway lost |] \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
307  | 
\ ==> Key K ~: analz (sees lost Spy evs)";  | 
| 2090 | 308  | 
by (forward_tac [Says_Server_message_form] 1 THEN assume_tac 1);  | 
| 2891 | 309  | 
by (blast_tac (!claset addSEs [lemma]) 1);  | 
| 2090 | 310  | 
qed "Spy_not_see_encrypted_key";  | 
311  | 
||
312  | 
||
313  | 
goal thy  | 
|
| 
2516
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
314  | 
 "!!evs. [| C ~: {A,B,Server};                                      \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
315  | 
\ Says Server B \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
316  | 
\              {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},    \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
317  | 
\                Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}   \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
318  | 
\ : set_of_list evs; \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
319  | 
\           Says B Spy {|NA, NB, Key K|} ~: set_of_list evs;        \
 | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
320  | 
\ A ~: lost; B ~: lost; evs : otway lost |] \  | 
| 
 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 
paulson 
parents: 
2454 
diff
changeset
 | 
321  | 
\ ==> Key K ~: analz (sees lost C evs)";  | 
| 2090 | 322  | 
by (rtac (subset_insertI RS sees_mono RS analz_mono RS contra_subsetD) 1);  | 
323  | 
by (rtac (sees_lost_agent_subset_sees_Spy RS analz_mono RS contra_subsetD) 1);  | 
|
324  | 
by (FIRSTGOAL (rtac Spy_not_see_encrypted_key));  | 
|
| 2891 | 325  | 
by (REPEAT_FIRST (blast_tac (!claset addIs [otway_mono RS subsetD])));  | 
| 2090 | 326  | 
qed "Agent_not_see_encrypted_key";  | 
327  | 
||
328  | 
||
329  | 
(**** Authenticity properties relating to NB ****)  | 
|
330  | 
||
331  | 
(*If the encrypted message appears then it originated with the Server!*)  | 
|
332  | 
goal thy  | 
|
| 2106 | 333  | 
"!!evs. [| B ~: lost; evs : otway lost |] \  | 
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
334  | 
\    ==> Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}                     \
 | 
| 2106 | 335  | 
\ : parts (sees lost Spy evs) \  | 
| 2090 | 336  | 
\ --> (EX NA. Says Server B \  | 
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
337  | 
\                     {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},     \
 | 
| 
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
338  | 
\                       Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}    \
 | 
| 2090 | 339  | 
\ : set_of_list evs)";  | 
| 
3121
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
340  | 
by parts_induct_tac;  | 
| 
 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 
paulson 
parents: 
3102 
diff
changeset
 | 
341  | 
by (Fake_parts_insert_tac 1);  | 
| 2090 | 342  | 
by (ALLGOALS (asm_simp_tac (!simpset addsimps [ex_disj_distrib])));  | 
343  | 
(*OR3*)  | 
|
| 3102 | 344  | 
by (Blast_tac 1);  | 
| 2090 | 345  | 
qed_spec_mp "NB_Crypt_imp_Server_msg";  | 
346  | 
||
347  | 
||
| 2454 | 348  | 
(*Guarantee for B: if it gets a well-formed certificate then the Server  | 
349  | 
has sent the correct message in round 3.*)  | 
|
| 2090 | 350  | 
goal thy  | 
| 2106 | 351  | 
"!!evs. [| B ~: lost; evs : otway lost; \  | 
| 
2837
 
dee1c7f1f576
Trivial renamings (for consistency with CSFW papers)
 
paulson 
parents: 
2637 
diff
changeset
 | 
352  | 
\           Says S' B {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} \
 | 
| 2106 | 353  | 
\ : set_of_list evs |] \  | 
354  | 
\ ==> EX NA. Says Server B \  | 
|
| 
2284
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
355  | 
\                     {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},     \
 | 
| 
 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 
paulson 
parents: 
2264 
diff
changeset
 | 
356  | 
\                       Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}    \
 | 
| 2090 | 357  | 
\ : set_of_list evs";  | 
| 3102 | 358  | 
by (blast_tac (!claset addSIs [NB_Crypt_imp_Server_msg]  | 
359  | 
addEs sees_Spy_partsEs) 1);  | 
|
| 2331 | 360  | 
qed "B_trusts_OR3";  |