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