| author | paulson | 
| Thu, 20 Nov 1997 11:55:39 +0100 | |
| changeset 4246 | c539e702e1d2 | 
| parent 4237 | fb01353e363b | 
| child 4267 | cdc193e38925 | 
| permissions | -rw-r--r-- | 
| 1934 | 1 | (* Title: HOL/Auth/NS_Shared | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | ||
| 6 | Inductive relation "ns_shared" for Needham-Schroeder Shared-Key protocol. | |
| 7 | ||
| 8 | From page 247 of | |
| 9 | Burrows, Abadi and Needham. A Logic of Authentication. | |
| 10 | Proc. Royal Soc. 426 (1989) | |
| 11 | *) | |
| 12 | ||
| 13 | open NS_Shared; | |
| 14 | ||
| 1943 | 15 | proof_timing:=true; | 
| 1997 | 16 | HOL_quantifiers := false; | 
| 17 | ||
| 18 | ||
| 2323 | 19 | (*A "possibility property": there are traces that reach the end*) | 
| 1997 | 20 | goal thy | 
| 3466 
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
 paulson parents: 
3465diff
changeset | 21 | "!!A B. [| A ~= B; A ~= Server; B ~= Server |] \ | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 22 | \ ==> EX N K. EX evs: ns_shared. \ | 
| 3465 | 23 | \               Says A B (Crypt K {|Nonce N, Nonce N|}) : set evs";
 | 
| 1997 | 24 | by (REPEAT (resolve_tac [exI,bexI] 1)); | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 25 | by (rtac (ns_shared.Nil RS ns_shared.NS1 RS ns_shared.NS2 RS | 
| 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 26 | ns_shared.NS3 RS ns_shared.NS4 RS ns_shared.NS5) 2); | 
| 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 27 | by possibility_tac; | 
| 2015 | 28 | result(); | 
| 29 | ||
| 3675 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 30 | goal thy | 
| 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 31 | "!!A B. [| A ~= B; A ~= Server; B ~= Server |] \ | 
| 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 32 | \ ==> EX evs: ns_shared. \ | 
| 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 33 | \               Says A B (Crypt ?K {|Nonce ?N, Nonce ?N|}) : set evs";
 | 
| 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 34 | by (REPEAT (resolve_tac [exI,bexI] 1)); | 
| 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 35 | by (rtac (ns_shared.Nil RS ns_shared.NS1 RS ns_shared.NS2 RS | 
| 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 36 | ns_shared.NS3 RS ns_shared.NS4 RS ns_shared.NS5) 2); | 
| 
70dd312b70b2
Deleted the redundant simprule not_parts_not_analz
 paulson parents: 
3651diff
changeset | 37 | by possibility_tac; | 
| 1943 | 38 | |
| 1934 | 39 | (**** Inductive proofs about ns_shared ****) | 
| 40 | ||
| 41 | (*Nobody sends themselves messages*) | |
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 42 | goal thy "!!evs. evs : ns_shared ==> ALL A X. Says A A X ~: set evs"; | 
| 2032 | 43 | by (etac ns_shared.induct 1); | 
| 1934 | 44 | by (Auto_tac()); | 
| 45 | qed_spec_mp "not_Says_to_self"; | |
| 46 | Addsimps [not_Says_to_self]; | |
| 47 | AddSEs [not_Says_to_self RSN (2, rev_notE)]; | |
| 48 | ||
| 1943 | 49 | (*For reasoning about the encrypted portion of message NS3*) | 
| 3465 | 50 | goal thy "!!evs. Says S A (Crypt KA {|N, B, K, X|}) : set evs \
 | 
| 3683 | 51 | \ ==> X : parts (spies evs)"; | 
| 4091 | 52 | by (blast_tac (claset() addSEs spies_partsEs) 1); | 
| 3683 | 53 | qed "NS3_msg_in_parts_spies"; | 
| 2032 | 54 | |
| 2070 | 55 | goal thy | 
| 3465 | 56 |     "!!evs. Says Server A (Crypt (shrK A) {|NA, B, K, X|}) : set evs \
 | 
| 3683 | 57 | \ ==> K : parts (spies evs)"; | 
| 4091 | 58 | by (blast_tac (claset() addSEs spies_partsEs) 1); | 
| 3683 | 59 | qed "Oops_parts_spies"; | 
| 2070 | 60 | |
| 3683 | 61 | (*For proving the easier theorems about X ~: parts (spies evs).*) | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 62 | fun parts_induct_tac i = | 
| 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 63 | etac ns_shared.induct i THEN | 
| 3683 | 64 | forward_tac [Oops_parts_spies] (i+7) THEN | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 65 | forward_tac [NS3_msg_in_parts_spies] (i+4) THEN | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 66 | prove_simple_subgoals_tac i; | 
| 2070 | 67 | |
| 1934 | 68 | |
| 3683 | 69 | (** Theorems of the form X ~: parts (spies evs) imply that NOBODY | 
| 2015 | 70 | sends messages containing X! **) | 
| 1934 | 71 | |
| 3683 | 72 | (*Spy never sees another agent's shared key! (unless it's bad at start)*) | 
| 1934 | 73 | goal thy | 
| 3683 | 74 | "!!evs. evs : ns_shared ==> (Key (shrK A) : parts (spies evs)) = (A : bad)"; | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 75 | by (parts_induct_tac 1); | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 76 | by (Fake_parts_insert_tac 1); | 
| 3961 | 77 | by (ALLGOALS Blast_tac); | 
| 2131 | 78 | qed "Spy_see_shrK"; | 
| 79 | Addsimps [Spy_see_shrK]; | |
| 1934 | 80 | |
| 2131 | 81 | goal thy | 
| 3683 | 82 | "!!evs. evs : ns_shared ==> (Key (shrK A) : analz (spies evs)) = (A : bad)"; | 
| 4091 | 83 | by (auto_tac(claset() addDs [impOfSubs analz_subset_parts], simpset())); | 
| 2131 | 84 | qed "Spy_analz_shrK"; | 
| 85 | Addsimps [Spy_analz_shrK]; | |
| 1934 | 86 | |
| 3683 | 87 | goal thy "!!A. [| Key (shrK A) : parts (spies evs); \ | 
| 88 | \ evs : ns_shared |] ==> A:bad"; | |
| 4091 | 89 | by (blast_tac (claset() addDs [Spy_see_shrK]) 1); | 
| 2131 | 90 | qed "Spy_see_shrK_D"; | 
| 1934 | 91 | |
| 2131 | 92 | bind_thm ("Spy_analz_shrK_D", analz_subset_parts RS subsetD RS Spy_see_shrK_D);
 | 
| 93 | AddSDs [Spy_see_shrK_D, Spy_analz_shrK_D]; | |
| 1934 | 94 | |
| 2070 | 95 | |
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 96 | (*Nobody can have used non-existent keys!*) | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 97 | goal thy "!!evs. evs : ns_shared ==> \ | 
| 3683 | 98 | \ Key K ~: used evs --> K ~: keysFor (parts (spies evs))"; | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 99 | by (parts_induct_tac 1); | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 100 | (*Fake*) | 
| 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 101 | by (best_tac | 
| 4091 | 102 | (claset() addSDs [impOfSubs (parts_insert_subset_Un RS keysFor_mono)] | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 103 | addIs [impOfSubs analz_subset_parts] | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 104 | addDs [impOfSubs (analz_subset_parts RS keysFor_mono)] | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 105 | addss (simpset())) 1); | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 106 | (*NS2, NS4, NS5*) | 
| 4091 | 107 | by (ALLGOALS (blast_tac (claset() addSEs spies_partsEs))); | 
| 2160 | 108 | qed_spec_mp "new_keys_not_used"; | 
| 1934 | 109 | |
| 110 | bind_thm ("new_keys_not_analzd",
 | |
| 2032 | 111 | [analz_subset_parts RS keysFor_mono, | 
| 112 | new_keys_not_used] MRS contra_subsetD); | |
| 1934 | 113 | |
| 114 | Addsimps [new_keys_not_used, new_keys_not_analzd]; | |
| 115 | ||
| 116 | ||
| 117 | (** Lemmas concerning the form of items passed in messages **) | |
| 118 | ||
| 2015 | 119 | (*Describes the form of K, X and K' when the Server sends this message.*) | 
| 1934 | 120 | goal thy | 
| 3683 | 121 |  "!!evs. [| Says Server A (Crypt K' {|N, Agent B, Key K, X|}) : set evs; \
 | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 122 | \ evs : ns_shared |] \ | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 123 | \ ==> K ~: range shrK & \ | 
| 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 124 | \            X = (Crypt (shrK B) {|Key K, Agent A|}) &        \
 | 
| 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 125 | \ K' = shrK A"; | 
| 2032 | 126 | by (etac rev_mp 1); | 
| 127 | by (etac ns_shared.induct 1); | |
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 128 | by (Auto_tac()); | 
| 1934 | 129 | qed "Says_Server_message_form"; | 
| 130 | ||
| 131 | ||
| 2070 | 132 | (*If the encrypted message appears then it originated with the Server*) | 
| 1934 | 133 | goal thy | 
| 3683 | 134 |  "!!evs. [| Crypt (shrK A) {|NA, Agent B, Key K, X|} : parts (spies evs); \
 | 
| 135 | \ A ~: bad; evs : ns_shared |] \ | |
| 3651 | 136 | \         ==> Says Server A (Crypt (shrK A) {|NA, Agent B, Key K, X|})  \
 | 
| 137 | \ : set evs"; | |
| 2070 | 138 | by (etac rev_mp 1); | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 139 | by (parts_induct_tac 1); | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 140 | by (Fake_parts_insert_tac 1); | 
| 2323 | 141 | qed "A_trusts_NS2"; | 
| 1934 | 142 | |
| 1965 | 143 | |
| 144 | (*EITHER describes the form of X when the following message is sent, | |
| 145 | OR reduces it to the Fake case. | |
| 146 | Use Says_Server_message_form if applicable.*) | |
| 1934 | 147 | goal thy | 
| 2284 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2264diff
changeset | 148 |  "!!evs. [| Says S A (Crypt (shrK A) {|Nonce NA, Agent B, Key K, X|})      \
 | 
| 3651 | 149 | \ : set evs; \ | 
| 150 | \ evs : ns_shared |] \ | |
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 151 | \        ==> (K ~: range shrK & X = (Crypt (shrK B) {|Key K, Agent A|}))   \
 | 
| 3683 | 152 | \ | X : analz (spies evs)"; | 
| 153 | by (case_tac "A : bad" 1); | |
| 4091 | 154 | by (fast_tac (claset() addSDs [Says_imp_spies RS analz.Inj] | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 155 | addss (simpset())) 1); | 
| 3683 | 156 | by (forward_tac [Says_imp_spies RS parts.Inj] 1); | 
| 4091 | 157 | by (blast_tac (claset() addSDs [A_trusts_NS2, Says_Server_message_form]) 1); | 
| 1934 | 158 | qed "Says_S_message_form"; | 
| 159 | ||
| 160 | ||
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 161 | (*For proofs involving analz.*) | 
| 3683 | 162 | val analz_spies_tac = | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 163 | forward_tac [Says_Server_message_form] 8 THEN | 
| 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 164 | forward_tac [Says_S_message_form] 5 THEN | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 165 | REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE] ORELSE' hyp_subst_tac); | 
| 2131 | 166 | |
| 1934 | 167 | |
| 168 | (**** | |
| 169 | The following is to prove theorems of the form | |
| 170 | ||
| 3683 | 171 | Key K : analz (insert (Key KAB) (spies evs)) ==> | 
| 172 | Key K : analz (spies evs) | |
| 1934 | 173 | |
| 174 | A more general formula must be proved inductively. | |
| 175 | ||
| 176 | ****) | |
| 177 | ||
| 178 | ||
| 179 | (*NOT useful in this form, but it says that session keys are not used | |
| 180 | to encrypt messages containing other keys, in the actual protocol. | |
| 181 | We require that agents should behave like this subsequently also.*) | |
| 182 | goal thy | |
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 183 | "!!evs. [| evs : ns_shared; Kab ~: range shrK |] ==> \ | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 184 | \ (Crypt KAB X) : parts (spies evs) & \ | 
| 3683 | 185 | \           Key K : parts {X} --> Key K : parts (spies evs)";
 | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 186 | by (parts_induct_tac 1); | 
| 1934 | 187 | (*Deals with Faked messages*) | 
| 4091 | 188 | by (blast_tac (claset() addSEs partsEs | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 189 | addDs [impOfSubs parts_insert_subset_Un]) 1); | 
| 1965 | 190 | (*Base, NS4 and NS5*) | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 191 | by (Auto_tac()); | 
| 1934 | 192 | result(); | 
| 193 | ||
| 194 | ||
| 195 | (** Session keys are not used to encrypt other session keys **) | |
| 196 | ||
| 2015 | 197 | (*The equality makes the induction hypothesis easier to apply*) | 
| 1934 | 198 | goal thy | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 199 | "!!evs. evs : ns_shared ==> \ | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 200 | \ ALL K KK. KK <= Compl (range shrK) --> \ | 
| 3683 | 201 | \ (Key K : analz (Key``KK Un (spies evs))) = \ | 
| 202 | \ (K : KK | Key K : analz (spies evs))"; | |
| 2032 | 203 | by (etac ns_shared.induct 1); | 
| 3683 | 204 | by analz_spies_tac; | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 205 | by (REPEAT_FIRST (resolve_tac [allI, impI])); | 
| 3961 | 206 | by (REPEAT_FIRST (rtac analz_image_freshK_lemma)); | 
| 207 | (*Takes 9 secs*) | |
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 208 | by (ALLGOALS (asm_simp_tac analz_image_freshK_ss)); | 
| 3451 
d10f100676d8
Made proofs more concise by replacing calls to spy_analz_tac by uses of
 paulson parents: 
3441diff
changeset | 209 | (*Fake*) | 
| 
d10f100676d8
Made proofs more concise by replacing calls to spy_analz_tac by uses of
 paulson parents: 
3441diff
changeset | 210 | by (spy_analz_tac 2); | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 211 | (*Base*) | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 212 | by (Blast_tac 1); | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 213 | qed_spec_mp "analz_image_freshK"; | 
| 1934 | 214 | |
| 215 | ||
| 216 | goal thy | |
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 217 | "!!evs. [| evs : ns_shared; KAB ~: range shrK |] ==> \ | 
| 3683 | 218 | \ Key K : analz (insert (Key KAB) (spies evs)) = \ | 
| 219 | \ (K = KAB | Key K : analz (spies evs))"; | |
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 220 | 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: 
2451diff
changeset | 221 | qed "analz_insert_freshK"; | 
| 1934 | 222 | |
| 223 | ||
| 2558 | 224 | (** The session key K uniquely identifies the message **) | 
| 1965 | 225 | |
| 1934 | 226 | goal thy | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 227 | "!!evs. evs : ns_shared ==> \ | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 228 | \ EX A' NA' B' X'. ALL A NA B X. \ | 
| 3683 | 229 | \       Says Server A (Crypt (shrK A) {|NA, Agent B, Key K, X|}) : set evs \
 | 
| 230 | \ --> A=A' & NA=NA' & B=B' & X=X'"; | |
| 2032 | 231 | by (etac ns_shared.induct 1); | 
| 4091 | 232 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib]))); | 
| 3730 | 233 | by Safe_tac; | 
| 2070 | 234 | (*NS3*) | 
| 235 | by (ex_strip_tac 2); | |
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 236 | by (Blast_tac 2); | 
| 2070 | 237 | (*NS2: it can't be a new key*) | 
| 238 | by (expand_case_tac "K = ?y" 1); | |
| 239 | by (REPEAT (ares_tac [refl,exI,impI,conjI] 2)); | |
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 240 | by (blast_tac (claset() delrules [conjI] (*prevent splitup into 4 subgoals*) | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 241 | addSEs spies_partsEs) 1); | 
| 1934 | 242 | val lemma = result(); | 
| 243 | ||
| 244 | (*In messages of this form, the session key uniquely identifies the rest*) | |
| 245 | goal thy | |
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 246 | "!!evs. [| Says Server A \ | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 247 | \             (Crypt (shrK A) {|NA, Agent B, Key K, X|}) : set evs;     \ 
 | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 248 | \ Says Server A' \ | 
| 3683 | 249 | \             (Crypt (shrK A') {|NA', Agent B', Key K, X'|}) : set evs; \
 | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 250 | \ evs : ns_shared |] ==> A=A' & NA=NA' & B=B' & X = X'"; | 
| 2451 
ce85a2aafc7a
Extensive tidying and simplification, largely stemming from
 paulson parents: 
2374diff
changeset | 251 | by (prove_unique_tac lemma 1); | 
| 1934 | 252 | qed "unique_session_keys"; | 
| 253 | ||
| 254 | ||
| 2032 | 255 | (** Crucial secrecy property: Spy does not see the keys sent in msg NS2 **) | 
| 2015 | 256 | |
| 1934 | 257 | goal thy | 
| 3961 | 258 | "!!evs. [| A ~: bad; B ~: bad; evs : ns_shared |] \ | 
| 2015 | 259 | \ ==> Says Server A \ | 
| 2284 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2264diff
changeset | 260 | \              (Crypt (shrK A) {|NA, Agent B, Key K,                   \
 | 
| 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2264diff
changeset | 261 | \                                Crypt (shrK B) {|Key K, Agent A|}|})  \
 | 
| 3466 
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
 paulson parents: 
3465diff
changeset | 262 | \ : set evs --> \ | 
| 
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
 paulson parents: 
3465diff
changeset | 263 | \        (ALL NB. Says A Spy {|NA, NB, Key K|} ~: set evs) -->         \
 | 
| 3683 | 264 | \ Key K ~: analz (spies evs)"; | 
| 2032 | 265 | by (etac ns_shared.induct 1); | 
| 3683 | 266 | by analz_spies_tac; | 
| 1934 | 267 | by (ALLGOALS | 
| 2015 | 268 | (asm_simp_tac | 
| 4091 | 269 | (simpset() addsimps ([analz_insert_eq, analz_insert_freshK] @ | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 270 | pushes @ expand_ifs)))); | 
| 3451 
d10f100676d8
Made proofs more concise by replacing calls to spy_analz_tac by uses of
 paulson parents: 
3441diff
changeset | 271 | (*Oops*) | 
| 4091 | 272 | by (blast_tac (claset() addDs [unique_session_keys]) 5); | 
| 3679 | 273 | (*NS3, replay sub-case*) | 
| 3451 
d10f100676d8
Made proofs more concise by replacing calls to spy_analz_tac by uses of
 paulson parents: 
3441diff
changeset | 274 | by (Blast_tac 4); | 
| 1934 | 275 | (*NS2*) | 
| 4091 | 276 | by (blast_tac (claset() addSEs spies_partsEs | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 277 | addIs [parts_insertI, | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 278 | impOfSubs analz_subset_parts]) 2); | 
| 3451 
d10f100676d8
Made proofs more concise by replacing calls to spy_analz_tac by uses of
 paulson parents: 
3441diff
changeset | 279 | (*Fake*) | 
| 
d10f100676d8
Made proofs more concise by replacing calls to spy_analz_tac by uses of
 paulson parents: 
3441diff
changeset | 280 | by (spy_analz_tac 1); | 
| 3679 | 281 | (*NS3, Server sub-case*) (**LEVEL 6 **) | 
| 4091 | 282 | by (clarify_tac (claset() delrules [impCE]) 1); | 
| 3683 | 283 | by (forward_tac [Says_imp_spies RS parts.Inj RS A_trusts_NS2] 1); | 
| 2170 | 284 | by (assume_tac 2); | 
| 4091 | 285 | by (blast_tac (claset() addDs [Says_imp_spies RS analz.Inj RS | 
| 3683 | 286 | Crypt_Spy_analz_bad]) 1); | 
| 4091 | 287 | by (blast_tac (claset() addDs [unique_session_keys]) 1); | 
| 2070 | 288 | val lemma = result() RS mp RS mp RSN(2,rev_notE); | 
| 2015 | 289 | |
| 290 | ||
| 291 | (*Final version: Server's message in the most abstract form*) | |
| 292 | goal thy | |
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 293 | "!!evs. [| Says Server A \ | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 294 | \            (Crypt K' {|NA, Agent B, Key K, X|}) : set evs;     \
 | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 295 | \           (ALL NB. Says A Spy {|NA, NB, Key K|} ~: set evs);   \
 | 
| 3683 | 296 | \ A ~: bad; B ~: bad; evs : ns_shared \ | 
| 297 | \ |] ==> Key K ~: analz (spies evs)"; | |
| 2015 | 298 | by (forward_tac [Says_Server_message_form] 1 THEN assume_tac 1); | 
| 4091 | 299 | by (blast_tac (claset() addSDs [lemma]) 1); | 
| 2032 | 300 | qed "Spy_not_see_encrypted_key"; | 
| 301 | ||
| 302 | ||
| 2070 | 303 | (**** Guarantees available at various stages of protocol ***) | 
| 304 | ||
| 3651 | 305 | A_trusts_NS2 RS Spy_not_see_encrypted_key; | 
| 2070 | 306 | |
| 307 | ||
| 308 | (*If the encrypted message appears then it originated with the Server*) | |
| 309 | goal thy | |
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 310 |  "!!evs. [| Crypt (shrK B) {|Key K, Agent A|} : parts (spies evs);     \
 | 
| 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 311 | \ B ~: bad; evs : ns_shared |] \ | 
| 2070 | 312 | \ ==> EX NA. Says Server A \ | 
| 2284 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2264diff
changeset | 313 | \              (Crypt (shrK A) {|NA, Agent B, Key K,                   \
 | 
| 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2264diff
changeset | 314 | \                                Crypt (shrK B) {|Key K, Agent A|}|})  \
 | 
| 3465 | 315 | \ : set evs"; | 
| 2070 | 316 | by (etac rev_mp 1); | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3516diff
changeset | 317 | by (parts_induct_tac 1); | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 318 | by (Fake_parts_insert_tac 1); | 
| 2070 | 319 | (*Fake case*) | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 320 | by (ALLGOALS Blast_tac); | 
| 2323 | 321 | qed "B_trusts_NS3"; | 
| 2070 | 322 | |
| 323 | ||
| 324 | goal thy | |
| 3683 | 325 | "!!evs. [| B ~: bad; evs : ns_shared |] \ | 
| 326 | \ ==> Key K ~: analz (spies evs) --> \ | |
| 2284 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2264diff
changeset | 327 | \            Says Server A (Crypt (shrK A) {|NA, Agent B, Key K, X|})  \
 | 
| 3466 
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
 paulson parents: 
3465diff
changeset | 328 | \ : set evs --> \ | 
| 3683 | 329 | \ Crypt K (Nonce NB) : parts (spies evs) --> \ | 
| 3465 | 330 | \ Says B A (Crypt K (Nonce NB)) : set evs"; | 
| 2070 | 331 | by (etac ns_shared.induct 1); | 
| 332 | by (forward_tac [Says_S_message_form] 5 THEN assume_tac 5); | |
| 3683 | 333 | by (dtac NS3_msg_in_parts_spies 5); | 
| 334 | by (forward_tac [Oops_parts_spies] 8); | |
| 2103 | 335 | by (TRYALL (rtac impI)); | 
| 336 | by (REPEAT_FIRST | |
| 3683 | 337 | (dtac (spies_subset_spies_Says RS analz_mono RS contra_subsetD))); | 
| 4091 | 338 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib]))); | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 339 | (**LEVEL 7**) | 
| 4091 | 340 | by (blast_tac (claset() addSDs [Crypt_Fake_parts_insert] | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 341 | addDs [impOfSubs analz_subset_parts]) 1); | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 342 | by (Blast_tac 2); | 
| 2103 | 343 | by (REPEAT_FIRST (rtac impI ORELSE' etac conjE ORELSE' hyp_subst_tac)); | 
| 2516 
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
 paulson parents: 
2451diff
changeset | 344 | (*Subgoal 1: contradiction from the assumptions | 
| 3683 | 345 | Key K ~: used evsa and Crypt K (Nonce NB) : parts (spies evsa) *) | 
| 2170 | 346 | by (dtac Crypt_imp_invKey_keysFor 1); | 
| 3121 
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
 paulson parents: 
2637diff
changeset | 347 | (**LEVEL 11**) | 
| 2070 | 348 | by (Asm_full_simp_tac 1); | 
| 2170 | 349 | by (rtac disjI1 1); | 
| 2070 | 350 | by (thin_tac "?PP-->?QQ" 1); | 
| 3683 | 351 | by (case_tac "Ba : bad" 1); | 
| 4091 | 352 | by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj RS B_trusts_NS3, | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 353 | unique_session_keys]) 2); | 
| 4091 | 354 | by (blast_tac (claset() addDs [Says_imp_spies RS analz.Inj RS | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 355 | Crypt_Spy_analz_bad]) 1); | 
| 2103 | 356 | val lemma = result(); | 
| 357 | ||
| 358 | goal thy | |
| 3683 | 359 | "!!evs. [| Crypt K (Nonce NB) : parts (spies evs); \ | 
| 2284 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2264diff
changeset | 360 | \           Says Server A (Crypt (shrK A) {|NA, Agent B, Key K, X|})  \
 | 
| 3466 
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
 paulson parents: 
3465diff
changeset | 361 | \ : set evs; \ | 
| 
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
 paulson parents: 
3465diff
changeset | 362 | \           ALL NB. Says A Spy {|NA, NB, Key K|} ~: set evs;          \
 | 
| 3683 | 363 | \ A ~: bad; B ~: bad; evs : ns_shared |] \ | 
| 3465 | 364 | \ ==> Says B A (Crypt K (Nonce NB)) : set evs"; | 
| 4091 | 365 | by (blast_tac (claset() addSIs [lemma RS mp RS mp RS mp] | 
| 4237 
fb01353e363b
The dtac was discarding information, though apparently no proofs were hurt
 paulson parents: 
4091diff
changeset | 366 | addSEs [Spy_not_see_encrypted_key RSN (2,rev_notE)]) 1); | 
| 2323 | 367 | qed "A_trusts_NS4"; |