| author | haftmann | 
| Wed, 15 Feb 2006 17:09:45 +0100 | |
| changeset 19042 | 630b8dd0b31a | 
| parent 18886 | 9f27383426db | 
| child 23746 | a455e69c31cc | 
| permissions | -rw-r--r-- | 
| 1934 | 1 | (* Title: HOL/Auth/NS_Shared | 
| 2 | ID: $Id$ | |
| 18886 | 3 | Author: Lawrence C Paulson and Giampaolo Bella | 
| 1934 | 4 | Copyright 1996 University of Cambridge | 
| 5 | *) | |
| 6 | ||
| 18886 | 7 | header{*Needham-Schroeder Shared-Key Protocol and the Issues Property*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 8 | |
| 16417 | 9 | theory NS_Shared imports Public begin | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 10 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 11 | text{*
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 12 | From page 247 of | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 13 | Burrows, Abadi and Needham (1989). A Logic of Authentication. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 14 | Proc. Royal Soc. 426 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 15 | *} | 
| 1934 | 16 | |
| 18886 | 17 | constdefs | 
| 18 | ||
| 19 | (* A is the true creator of X if she has sent X and X never appeared on | |
| 20 | the trace before this event. Recall that traces grow from head. *) | |
| 21 | Issues :: "[agent, agent, msg, event list] => bool" | |
| 22 |              ("_ Issues _ with _ on _")
 | |
| 23 | "A Issues B with X on evs == | |
| 24 |       \<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
 | |
| 25 | X \<notin> parts (spies (takeWhile (% z. z \<noteq> Says A B Y) (rev evs)))" | |
| 26 | ||
| 27 | ||
| 11104 | 28 | consts ns_shared :: "event list set" | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3465diff
changeset | 29 | inductive "ns_shared" | 
| 11104 | 30 | intros | 
| 31 | (*Initial trace is empty*) | |
| 13926 | 32 | Nil: "[] \<in> ns_shared" | 
| 11104 | 33 | (*The spy MAY say anything he CAN say. We do not expect him to | 
| 34 | invent new nonces here, but he can also use NS1. Common to | |
| 35 | all similar protocols.*) | |
| 13926 | 36 | Fake: "\<lbrakk>evsf \<in> ns_shared; X \<in> synth (analz (spies evsf))\<rbrakk> | 
| 37 | \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared" | |
| 11104 | 38 | |
| 39 | (*Alice initiates a protocol run, requesting to talk to any B*) | |
| 13926 | 40 | NS1: "\<lbrakk>evs1 \<in> ns_shared; Nonce NA \<notin> used evs1\<rbrakk> | 
| 41 | \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared" | |
| 11104 | 42 | |
| 43 | (*Server's response to Alice's message. | |
| 44 | !! It may respond more than once to A's request !! | |
| 45 | Server doesn't know who the true sender is, hence the A' in | |
| 46 | the sender field.*) | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 47 | NS2: "\<lbrakk>evs2 \<in> ns_shared; Key KAB \<notin> used evs2; KAB \<in> symKeys; | 
| 13926 | 48 | Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk> | 
| 49 | \<Longrightarrow> Says Server A | |
| 11104 | 50 | (Crypt (shrK A) | 
| 13926 | 51 | \<lbrace>Nonce NA, Agent B, Key KAB, | 
| 52 | (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>) | |
| 53 | # evs2 \<in> ns_shared" | |
| 11104 | 54 | |
| 55 | (*We can't assume S=Server. Agent A "remembers" her nonce. | |
| 13926 | 56 | Need A \<noteq> Server because we allow messages to self.*) | 
| 57 | NS3: "\<lbrakk>evs3 \<in> ns_shared; A \<noteq> Server; | |
| 58 | Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3; | |
| 59 | Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk> | |
| 60 | \<Longrightarrow> Says A B X # evs3 \<in> ns_shared" | |
| 11104 | 61 | |
| 62 | (*Bob's nonce exchange. He does not know who the message came | |
| 63 | from, but responds to A because she is mentioned inside.*) | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 64 | NS4: "\<lbrakk>evs4 \<in> ns_shared; Nonce NB \<notin> used evs4; K \<in> symKeys; | 
| 13926 | 65 | Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk> | 
| 66 | \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared" | |
| 1934 | 67 | |
| 11104 | 68 | (*Alice responds with Nonce NB if she has seen the key before. | 
| 69 | Maybe should somehow check Nonce NA again. | |
| 70 | We do NOT send NB-1 or similar as the Spy cannot spoof such things. | |
| 11465 | 71 | Letting the Spy add or subtract 1 lets him send all nonces. | 
| 11104 | 72 | Instead we distinguish the messages by sending the nonce twice.*) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 73 | NS5: "\<lbrakk>evs5 \<in> ns_shared; K \<in> symKeys; | 
| 13926 | 74 | Says B' A (Crypt K (Nonce NB)) \<in> set evs5; | 
| 75 | Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) | |
| 76 | \<in> set evs5\<rbrakk> | |
| 77 | \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared" | |
| 11104 | 78 | |
| 79 | (*This message models possible leaks of session keys. | |
| 80 | The two Nonces identify the protocol run: the rule insists upon | |
| 81 | the true senders in order to make them accurate.*) | |
| 13926 | 82 | Oops: "\<lbrakk>evso \<in> ns_shared; Says B A (Crypt K (Nonce NB)) \<in> set evso; | 
| 83 | Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) | |
| 84 | \<in> set evso\<rbrakk> | |
| 85 | \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared" | |
| 11104 | 86 | |
| 11150 | 87 | |
| 88 | declare Says_imp_knows_Spy [THEN parts.Inj, dest] | |
| 89 | declare parts.Body [dest] | |
| 11251 | 90 | declare Fake_parts_insert_in_Un [dest] | 
| 91 | declare analz_into_parts [dest] | |
| 11104 | 92 | declare image_eq_UN [simp] (*accelerates proofs involving nested images*) | 
| 93 | ||
| 94 | ||
| 13926 | 95 | text{*A "possibility property": there are traces that reach the end*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 96 | lemma "[| A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys |] | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
13956diff
changeset | 97 | ==> \<exists>N. \<exists>evs \<in> ns_shared. | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
13956diff
changeset | 98 | Says A B (Crypt K \<lbrace>Nonce N, Nonce N\<rbrace>) \<in> set evs" | 
| 11104 | 99 | apply (intro exI bexI) | 
| 100 | apply (rule_tac [2] ns_shared.Nil | |
| 101 | [THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3, | |
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
13956diff
changeset | 102 | THEN ns_shared.NS4, THEN ns_shared.NS5]) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 103 | apply (possibility, simp add: used_Cons) | 
| 11104 | 104 | done | 
| 105 | ||
| 106 | (*This version is similar, while instantiating ?K and ?N to epsilon-terms | |
| 13926 | 107 | lemma "A \<noteq> Server \<Longrightarrow> \<exists>evs \<in> ns_shared. | 
| 108 | Says A B (Crypt ?K \<lbrace>Nonce ?N, Nonce ?N\<rbrace>) \<in> set evs" | |
| 11104 | 109 | *) | 
| 110 | ||
| 111 | ||
| 13926 | 112 | subsection{*Inductive proofs about @{term ns_shared}*}
 | 
| 11104 | 113 | |
| 13926 | 114 | subsubsection{*Forwarding lemmas, to aid simplification*}
 | 
| 1934 | 115 | |
| 13926 | 116 | text{*For reasoning about the encrypted portion of message NS3*}
 | 
| 11104 | 117 | lemma NS3_msg_in_parts_spies: | 
| 13926 | 118 | "Says S A (Crypt KA \<lbrace>N, B, K, X\<rbrace>) \<in> set evs \<Longrightarrow> X \<in> parts (spies evs)" | 
| 11104 | 119 | by blast | 
| 11280 | 120 | |
| 13926 | 121 | text{*For reasoning about the Oops message*}
 | 
| 11104 | 122 | lemma Oops_parts_spies: | 
| 13926 | 123 | "Says Server A (Crypt (shrK A) \<lbrace>NA, B, K, X\<rbrace>) \<in> set evs | 
| 124 | \<Longrightarrow> K \<in> parts (spies evs)" | |
| 11104 | 125 | by blast | 
| 126 | ||
| 13926 | 127 | text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that NOBODY
 | 
| 128 |     sends messages containing @{term X}*}
 | |
| 11104 | 129 | |
| 13926 | 130 | text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
 | 
| 11104 | 131 | lemma Spy_see_shrK [simp]: | 
| 13926 | 132 | "evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" | 
| 13507 | 133 | apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all, blast+) | 
| 11104 | 134 | done | 
| 135 | ||
| 136 | lemma Spy_analz_shrK [simp]: | |
| 13926 | 137 | "evs \<in> ns_shared \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)" | 
| 11104 | 138 | by auto | 
| 139 | ||
| 140 | ||
| 13926 | 141 | text{*Nobody can have used non-existent keys!*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 142 | lemma new_keys_not_used [simp]: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 143 | "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> ns_shared|] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 144 | ==> K \<notin> keysFor (parts (spies evs))" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 145 | apply (erule rev_mp) | 
| 13507 | 146 | apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all) | 
| 13926 | 147 | txt{*Fake, NS2, NS4, NS5*}
 | 
| 148 | apply (force dest!: keysFor_parts_insert, blast+) | |
| 11104 | 149 | done | 
| 150 | ||
| 151 | ||
| 13926 | 152 | subsubsection{*Lemmas concerning the form of items passed in messages*}
 | 
| 11104 | 153 | |
| 13926 | 154 | text{*Describes the form of K, X and K' when the Server sends this message.*}
 | 
| 11104 | 155 | lemma Says_Server_message_form: | 
| 13926 | 156 | "\<lbrakk>Says Server A (Crypt K' \<lbrace>N, Agent B, Key K, X\<rbrace>) \<in> set evs; | 
| 157 | evs \<in> ns_shared\<rbrakk> | |
| 158 | \<Longrightarrow> K \<notin> range shrK \<and> | |
| 159 | X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<and> | |
| 11104 | 160 | K' = shrK A" | 
| 161 | by (erule rev_mp, erule ns_shared.induct, auto) | |
| 162 | ||
| 1934 | 163 | |
| 13926 | 164 | text{*If the encrypted message appears then it originated with the Server*}
 | 
| 11104 | 165 | lemma A_trusts_NS2: | 
| 13926 | 166 | "\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | 
| 167 | A \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 168 | \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs" | |
| 11104 | 169 | apply (erule rev_mp) | 
| 13507 | 170 | apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, auto) | 
| 11104 | 171 | done | 
| 172 | ||
| 173 | lemma cert_A_form: | |
| 13926 | 174 | "\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | 
| 175 | A \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 176 | \<Longrightarrow> K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)" | |
| 11104 | 177 | by (blast dest!: A_trusts_NS2 Says_Server_message_form) | 
| 178 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 179 | text{*EITHER describes the form of X when the following message is sent,
 | 
| 11104 | 180 | OR reduces it to the Fake case. | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 181 |   Use @{text Says_Server_message_form} if applicable.*}
 | 
| 11104 | 182 | lemma Says_S_message_form: | 
| 13926 | 183 | "\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs; | 
| 184 | evs \<in> ns_shared\<rbrakk> | |
| 185 | \<Longrightarrow> (K \<notin> range shrK \<and> X = (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>)) | |
| 186 | \<or> X \<in> analz (spies evs)" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 187 | by (blast dest: Says_imp_knows_Spy analz_shrK_Decrypt cert_A_form analz.Inj) | 
| 11150 | 188 | |
| 11104 | 189 | |
| 190 | (*Alternative version also provable | |
| 191 | lemma Says_S_message_form2: | |
| 13926 | 192 | "\<lbrakk>Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs; | 
| 193 | evs \<in> ns_shared\<rbrakk> | |
| 194 | \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs | |
| 195 | \<or> X \<in> analz (spies evs)" | |
| 196 | apply (case_tac "A \<in> bad") | |
| 13507 | 197 | apply (force dest!: Says_imp_knows_Spy [THEN analz.Inj]) | 
| 11104 | 198 | by (blast dest!: A_trusts_NS2 Says_Server_message_form) | 
| 199 | *) | |
| 200 | ||
| 201 | ||
| 202 | (**** | |
| 203 | SESSION KEY COMPROMISE THEOREM. To prove theorems of the form | |
| 204 | ||
| 13926 | 205 | Key K \<in> analz (insert (Key KAB) (spies evs)) \<Longrightarrow> | 
| 206 | Key K \<in> analz (spies evs) | |
| 11104 | 207 | |
| 208 | A more general formula must be proved inductively. | |
| 209 | ****) | |
| 1934 | 210 | |
| 13926 | 211 | text{*NOT useful in this form, but it says that session keys are not used
 | 
| 11104 | 212 | to encrypt messages containing other keys, in the actual protocol. | 
| 13926 | 213 | We require that agents should behave like this subsequently also.*} | 
| 214 | lemma "\<lbrakk>evs \<in> ns_shared; Kab \<notin> range shrK\<rbrakk> \<Longrightarrow> | |
| 215 | (Crypt KAB X) \<in> parts (spies evs) \<and> | |
| 216 |          Key K \<in> parts {X} \<longrightarrow> Key K \<in> parts (spies evs)"
 | |
| 13507 | 217 | apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, simp_all) | 
| 13926 | 218 | txt{*Fake*}
 | 
| 11104 | 219 | apply (blast dest: parts_insert_subset_Un) | 
| 13926 | 220 | txt{*Base, NS4 and NS5*}
 | 
| 11104 | 221 | apply auto | 
| 222 | done | |
| 223 | ||
| 224 | ||
| 13926 | 225 | subsubsection{*Session keys are not used to encrypt other session keys*}
 | 
| 11104 | 226 | |
| 13926 | 227 | text{*The equality makes the induction hypothesis easier to apply*}
 | 
| 11104 | 228 | |
| 229 | lemma analz_image_freshK [rule_format]: | |
| 13926 | 230 | "evs \<in> ns_shared \<Longrightarrow> | 
| 231 | \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow> | |
| 232 | (Key K \<in> analz (Key`KK \<union> (spies evs))) = | |
| 233 | (K \<in> KK \<or> Key K \<in> analz (spies evs))" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 234 | apply (erule ns_shared.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 235 | apply (drule_tac [8] Says_Server_message_form) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 236 | apply (erule_tac [5] Says_S_message_form [THEN disjE], analz_freshK, spy_analz) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 237 | txt{*NS2, NS3*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 238 | apply blast+; | 
| 11104 | 239 | done | 
| 240 | ||
| 241 | ||
| 242 | lemma analz_insert_freshK: | |
| 13926 | 243 | "\<lbrakk>evs \<in> ns_shared; KAB \<notin> range shrK\<rbrakk> \<Longrightarrow> | 
| 244 | (Key K \<in> analz (insert (Key KAB) (spies evs))) = | |
| 245 | (K = KAB \<or> Key K \<in> analz (spies evs))" | |
| 11104 | 246 | by (simp only: analz_image_freshK analz_image_freshK_simps) | 
| 247 | ||
| 248 | ||
| 13926 | 249 | subsubsection{*The session key K uniquely identifies the message*}
 | 
| 1934 | 250 | |
| 13926 | 251 | text{*In messages of this form, the session key uniquely identifies the rest*}
 | 
| 11104 | 252 | lemma unique_session_keys: | 
| 13926 | 253 | "\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs; | 
| 254 | Says Server A' (Crypt (shrK A') \<lbrace>NA', Agent B', Key K, X'\<rbrace>) \<in> set evs; | |
| 255 | evs \<in> ns_shared\<rbrakk> \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B' \<and> X = X'" | |
| 18886 | 256 | by (erule rev_mp, erule rev_mp, erule ns_shared.induct, simp_all, blast+) | 
| 11104 | 257 | |
| 258 | ||
| 18886 | 259 | subsubsection{*Crucial secrecy property: Spy doesn't see the keys sent in NS2*}
 | 
| 11104 | 260 | |
| 13956 | 261 | text{*Beware of @{text "[rule_format]"} and the universal quantifier!*}
 | 
| 11150 | 262 | lemma secrecy_lemma: | 
| 13926 | 263 | "\<lbrakk>Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, | 
| 264 | Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) | |
| 265 | \<in> set evs; | |
| 266 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 267 | \<Longrightarrow> (\<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs) \<longrightarrow> | |
| 268 | Key K \<notin> analz (spies evs)" | |
| 11104 | 269 | apply (erule rev_mp) | 
| 270 | apply (erule ns_shared.induct, force) | |
| 271 | apply (frule_tac [7] Says_Server_message_form) | |
| 272 | apply (frule_tac [4] Says_S_message_form) | |
| 273 | apply (erule_tac [5] disjE) | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 274 | apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs, spy_analz) | 
| 13926 | 275 | txt{*NS2*}
 | 
| 276 | apply blast | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 277 | txt{*NS3, Server sub-case*}
 | 
| 11188 | 278 | apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2 | 
| 279 | dest: Says_imp_knows_Spy analz.Inj unique_session_keys) | |
| 13926 | 280 | txt{*NS3, Spy sub-case; also Oops*}
 | 
| 11280 | 281 | apply (blast dest: unique_session_keys)+ | 
| 11104 | 282 | done | 
| 283 | ||
| 284 | ||
| 11188 | 285 | |
| 13926 | 286 | text{*Final version: Server's message in the most abstract form*}
 | 
| 11104 | 287 | lemma Spy_not_see_encrypted_key: | 
| 13926 | 288 | "\<lbrakk>Says Server A (Crypt K' \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs; | 
| 289 | \<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; | |
| 290 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 291 | \<Longrightarrow> Key K \<notin> analz (spies evs)" | |
| 11150 | 292 | by (blast dest: Says_Server_message_form secrecy_lemma) | 
| 11104 | 293 | |
| 294 | ||
| 13926 | 295 | subsection{*Guarantees available at various stages of protocol*}
 | 
| 1934 | 296 | |
| 13926 | 297 | text{*If the encrypted message appears then it originated with the Server*}
 | 
| 11104 | 298 | lemma B_trusts_NS3: | 
| 13926 | 299 | "\<lbrakk>Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs); | 
| 300 | B \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 301 | \<Longrightarrow> \<exists>NA. Says Server A | |
| 302 | (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, | |
| 303 | Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) | |
| 304 | \<in> set evs" | |
| 11104 | 305 | apply (erule rev_mp) | 
| 13507 | 306 | apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies, auto) | 
| 11104 | 307 | done | 
| 308 | ||
| 309 | ||
| 310 | lemma A_trusts_NS4_lemma [rule_format]: | |
| 13926 | 311 | "evs \<in> ns_shared \<Longrightarrow> | 
| 312 | Key K \<notin> analz (spies evs) \<longrightarrow> | |
| 313 | Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow> | |
| 314 | Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow> | |
| 315 | Says B A (Crypt K (Nonce NB)) \<in> set evs" | |
| 11104 | 316 | apply (erule ns_shared.induct, force, drule_tac [4] NS3_msg_in_parts_spies) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 317 | apply (analz_mono_contra, simp_all, blast) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 318 | txt{*NS2: contradiction from the assumptions @{term "Key K \<notin> used evs2"} and
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 319 |     @{term "Crypt K (Nonce NB) \<in> parts (spies evs2)"} *} 
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 320 | apply (force dest!: Crypt_imp_keysFor) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 321 | txt{*NS4*}
 | 
| 13935 | 322 | apply (blast dest: B_trusts_NS3 | 
| 323 | Says_imp_knows_Spy [THEN analz.Inj] | |
| 11150 | 324 | Crypt_Spy_analz_bad unique_session_keys) | 
| 11104 | 325 | done | 
| 326 | ||
| 13926 | 327 | text{*This version no longer assumes that K is secure*}
 | 
| 11104 | 328 | lemma A_trusts_NS4: | 
| 13926 | 329 | "\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs); | 
| 330 | Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | |
| 331 | \<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; | |
| 332 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 333 | \<Longrightarrow> Says B A (Crypt K (Nonce NB)) \<in> set evs" | |
| 11280 | 334 | by (blast intro: A_trusts_NS4_lemma | 
| 11104 | 335 | dest: A_trusts_NS2 Spy_not_see_encrypted_key) | 
| 336 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 337 | text{*If the session key has been used in NS4 then somebody has forwarded
 | 
| 11280 | 338 | component X in some instance of NS4. Perhaps an interesting property, | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 339 | but not needed (after all) for the proofs below.*} | 
| 11104 | 340 | theorem NS4_implies_NS3 [rule_format]: | 
| 13926 | 341 | "evs \<in> ns_shared \<Longrightarrow> | 
| 342 | Key K \<notin> analz (spies evs) \<longrightarrow> | |
| 343 | Says Server A (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace>) \<in> set evs \<longrightarrow> | |
| 344 | Crypt K (Nonce NB) \<in> parts (spies evs) \<longrightarrow> | |
| 345 | (\<exists>A'. Says A' B X \<in> set evs)" | |
| 18886 | 346 | apply (erule ns_shared.induct, force) | 
| 347 | apply (drule_tac [4] NS3_msg_in_parts_spies) | |
| 348 | apply analz_mono_contra | |
| 13926 | 349 | apply (simp_all add: ex_disj_distrib, blast) | 
| 350 | txt{*NS2*}
 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 351 | apply (blast dest!: new_keys_not_used Crypt_imp_keysFor) | 
| 13926 | 352 | txt{*NS4*}
 | 
| 13935 | 353 | apply (blast dest: B_trusts_NS3 | 
| 11280 | 354 | dest: Says_imp_knows_Spy [THEN analz.Inj] | 
| 11150 | 355 | unique_session_keys Crypt_Spy_analz_bad) | 
| 11104 | 356 | done | 
| 357 | ||
| 358 | ||
| 359 | lemma B_trusts_NS5_lemma [rule_format]: | |
| 13926 | 360 | "\<lbrakk>B \<notin> bad; evs \<in> ns_shared\<rbrakk> \<Longrightarrow> | 
| 361 | Key K \<notin> analz (spies evs) \<longrightarrow> | |
| 11104 | 362 | Says Server A | 
| 13926 | 363 | (Crypt (shrK A) \<lbrace>NA, Agent B, Key K, | 
| 364 | Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow> | |
| 365 | Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow> | |
| 366 | Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs" | |
| 18886 | 367 | apply (erule ns_shared.induct, force) | 
| 368 | apply (drule_tac [4] NS3_msg_in_parts_spies) | |
| 369 | apply (analz_mono_contra, simp_all, blast) | |
| 13926 | 370 | txt{*NS2*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 371 | apply (blast dest!: new_keys_not_used Crypt_imp_keysFor) | 
| 13926 | 372 | txt{*NS5*}
 | 
| 11150 | 373 | apply (blast dest!: A_trusts_NS2 | 
| 11280 | 374 | dest: Says_imp_knows_Spy [THEN analz.Inj] | 
| 11150 | 375 | unique_session_keys Crypt_Spy_analz_bad) | 
| 11104 | 376 | done | 
| 377 | ||
| 378 | ||
| 13926 | 379 | text{*Very strong Oops condition reveals protocol's weakness*}
 | 
| 11104 | 380 | lemma B_trusts_NS5: | 
| 13926 | 381 | "\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs); | 
| 382 | Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs); | |
| 383 | \<forall>NA NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; | |
| 384 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 385 | \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs" | |
| 11280 | 386 | by (blast intro: B_trusts_NS5_lemma | 
| 11150 | 387 | dest: B_trusts_NS3 Spy_not_see_encrypted_key) | 
| 1934 | 388 | |
| 18886 | 389 | text{*Unaltered so far wrt original version*}
 | 
| 390 | ||
| 391 | subsection{*Lemmas for reasoning about predicate "Issues"*}
 | |
| 392 | ||
| 393 | lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)" | |
| 394 | apply (induct_tac "evs") | |
| 395 | apply (induct_tac [2] "a", auto) | |
| 396 | done | |
| 397 | ||
| 398 | lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs" | |
| 399 | apply (induct_tac "evs") | |
| 400 | apply (induct_tac [2] "a", auto) | |
| 401 | done | |
| 402 | ||
| 403 | lemma spies_Notes_rev: "spies (evs @ [Notes A X]) = | |
| 404 | (if A:bad then insert X (spies evs) else spies evs)" | |
| 405 | apply (induct_tac "evs") | |
| 406 | apply (induct_tac [2] "a", auto) | |
| 407 | done | |
| 408 | ||
| 409 | lemma spies_evs_rev: "spies evs = spies (rev evs)" | |
| 410 | apply (induct_tac "evs") | |
| 411 | apply (induct_tac [2] "a") | |
| 412 | apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev) | |
| 413 | done | |
| 414 | ||
| 415 | lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono] | |
| 416 | ||
| 417 | lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs" | |
| 418 | apply (induct_tac "evs") | |
| 419 | apply (induct_tac [2] "a", auto) | |
| 420 | txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 | |
| 421 | done | |
| 422 | ||
| 423 | lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono] | |
| 424 | ||
| 425 | ||
| 426 | subsection{*Guarantees of non-injective agreement on the session key, and
 | |
| 427 | of key distribution. They also express forms of freshness of certain messages, | |
| 428 | namely that agents were alive after something happened.*} | |
| 429 | ||
| 430 | lemma B_Issues_A: | |
| 431 | "\<lbrakk> Says B A (Crypt K (Nonce Nb)) \<in> set evs; | |
| 432 | Key K \<notin> analz (spies evs); | |
| 433 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared \<rbrakk> | |
| 434 | \<Longrightarrow> B Issues A with (Crypt K (Nonce Nb)) on evs" | |
| 435 | apply (simp (no_asm) add: Issues_def) | |
| 436 | apply (rule exI) | |
| 437 | apply (rule conjI, assumption) | |
| 438 | apply (simp (no_asm)) | |
| 439 | apply (erule rev_mp) | |
| 440 | apply (erule rev_mp) | |
| 441 | apply (erule ns_shared.induct, analz_mono_contra) | |
| 442 | apply (simp_all) | |
| 443 | txt{*fake*}
 | |
| 444 | apply blast | |
| 445 | apply (simp_all add: takeWhile_tail) | |
| 446 | txt{*NS3 remains by pure coincidence!*}
 | |
| 447 | apply (force dest!: A_trusts_NS2 Says_Server_message_form) | |
| 448 | txt{*NS4 would be the non-trivial case can be solved by Nb being used*}
 | |
| 449 | apply (blast dest: parts_spies_takeWhile_mono [THEN subsetD] | |
| 450 | parts_spies_evs_revD2 [THEN subsetD]) | |
| 451 | done | |
| 452 | ||
| 453 | text{*Tells A that B was alive after she sent him the session key.  The
 | |
| 454 | session key must be assumed confidential for this deduction to be meaningful, | |
| 455 | but that assumption can be relaxed by the appropriate argument. | |
| 456 | ||
| 457 | Precisely, the theorem guarantees (to A) key distribution of the session key | |
| 458 | to B. It also guarantees (to A) non-injective agreement of B with A on the | |
| 459 | session key. Both goals are available to A in the sense of Goal Availability. | |
| 460 | *} | |
| 461 | lemma A_authenticates_and_keydist_to_B: | |
| 462 | "\<lbrakk>Crypt K (Nonce NB) \<in> parts (spies evs); | |
| 463 | Crypt (shrK A) \<lbrace>NA, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | |
| 464 | Key K \<notin> analz(knows Spy evs); | |
| 465 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 466 | \<Longrightarrow> B Issues A with (Crypt K (Nonce NB)) on evs" | |
| 467 | by (blast intro: A_trusts_NS4_lemma B_Issues_A dest: A_trusts_NS2) | |
| 468 | ||
| 469 | lemma A_trusts_NS5: | |
| 470 | "\<lbrakk> Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts(spies evs); | |
| 471 | Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace> \<in> parts(spies evs); | |
| 472 | Key K \<notin> analz (spies evs); | |
| 473 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared \<rbrakk> | |
| 474 | \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"; | |
| 475 | apply (erule rev_mp) | |
| 476 | apply (erule rev_mp) | |
| 477 | apply (erule rev_mp) | |
| 478 | apply (erule ns_shared.induct, analz_mono_contra) | |
| 479 | apply (frule_tac [5] Says_S_message_form) | |
| 480 | apply (simp_all) | |
| 481 | txt{*Fake*}
 | |
| 482 | apply blast | |
| 483 | txt{*NS2*}
 | |
| 484 | apply (force dest!: Crypt_imp_keysFor) | |
| 485 | txt{*NS3, much quicker having installed @{term Says_S_message_form} before simplication*}
 | |
| 486 | apply fastsimp | |
| 487 | txt{*NS5, the most important case, can be solved by unicity*}
 | |
| 488 | apply (case_tac "Aa \<in> bad") | |
| 489 | apply (force dest!: Says_imp_spies [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst]) | |
| 490 | apply (blast dest: A_trusts_NS2 unique_session_keys) | |
| 491 | done | |
| 492 | ||
| 493 | lemma A_Issues_B: | |
| 494 | "\<lbrakk> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs; | |
| 495 | Key K \<notin> analz (spies evs); | |
| 496 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared \<rbrakk> | |
| 497 | \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) on evs" | |
| 498 | apply (simp (no_asm) add: Issues_def) | |
| 499 | apply (rule exI) | |
| 500 | apply (rule conjI, assumption) | |
| 501 | apply (simp (no_asm)) | |
| 502 | apply (erule rev_mp) | |
| 503 | apply (erule rev_mp) | |
| 504 | apply (erule ns_shared.induct, analz_mono_contra) | |
| 505 | apply (simp_all) | |
| 506 | txt{*fake*}
 | |
| 507 | apply blast | |
| 508 | apply (simp_all add: takeWhile_tail) | |
| 509 | txt{*NS3 remains by pure coincidence!*}
 | |
| 510 | apply (force dest!: A_trusts_NS2 Says_Server_message_form) | |
| 511 | txt{*NS5 is the non-trivial case and cannot be solved as in @{term B_Issues_A}! because NB is not fresh. We need @{term A_trusts_NS5}, proved for this very purpose*}
 | |
| 512 | apply (blast dest: A_trusts_NS5 parts_spies_takeWhile_mono [THEN subsetD] | |
| 513 | parts_spies_evs_revD2 [THEN subsetD]) | |
| 514 | done | |
| 515 | ||
| 516 | text{*Tells B that A was alive after B issued NB.
 | |
| 517 | ||
| 518 | Precisely, the theorem guarantees (to B) key distribution of the session key to A. It also guarantees (to B) non-injective agreement of A with B on the session key. Both goals are available to B in the sense of Goal Availability. | |
| 519 | *} | |
| 520 | lemma B_authenticates_and_keydist_to_A: | |
| 521 | "\<lbrakk>Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs); | |
| 522 | Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace> \<in> parts (spies evs); | |
| 523 | Key K \<notin> analz (spies evs); | |
| 524 | A \<notin> bad; B \<notin> bad; evs \<in> ns_shared\<rbrakk> | |
| 525 | \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) on evs" | |
| 526 | by (blast intro: A_Issues_B B_trusts_NS5_lemma dest: B_trusts_NS3) | |
| 527 | ||
| 1934 | 528 | end |