| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 28 Jun 2024 11:09:04 +0200 | |
| changeset 80418 | 9f90c4864e55 | 
| parent 76287 | cdc14f94c754 | 
| permissions | -rw-r--r-- | 
| 37936 | 1 | (* Title: HOL/Auth/Yahalom_Bad.thy | 
| 6400 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1996 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 61830 | 6 | section\<open>The Yahalom Protocol: A Flawed Version\<close> | 
| 13956 | 7 | |
| 16417 | 8 | theory Yahalom_Bad 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 | Demonstrates of why Oops is necessary. This protocol can be attacked because | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 12 | it doesn't keep NB secret, but without Oops it can be "verified" anyway. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 13 | The issues are discussed in lcp's LICS 2000 invited lecture. | 
| 61830 | 14 | \<close> | 
| 6400 | 15 | |
| 23746 | 16 | inductive_set yahalom :: "event list set" | 
| 17 | where | |
| 6400 | 18 | (*Initial trace is empty*) | 
| 11251 | 19 | Nil: "[] \<in> yahalom" | 
| 6400 | 20 | |
| 21 | (*The spy MAY say anything he CAN say. We do not expect him to | |
| 22 | invent new nonces here, but he can also use NS1. Common to | |
| 23 | all similar protocols.*) | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 24 | | Fake: "\<lbrakk>evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf))\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 25 | \<Longrightarrow> Says Spy B X # evsf \<in> yahalom" | 
| 6400 | 26 | |
| 27 | (*A message that has been sent can be received by the | |
| 28 | intended recipient.*) | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 29 | | Reception: "\<lbrakk>evsr \<in> yahalom; Says A B X \<in> set evsr\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 30 | \<Longrightarrow> Gets B X # evsr \<in> yahalom" | 
| 6400 | 31 | |
| 32 | (*Alice initiates a protocol run*) | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 33 | | YM1: "\<lbrakk>evs1 \<in> yahalom; Nonce NA \<notin> used evs1\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 34 | \<Longrightarrow> Says A B \<lbrace>Agent A, Nonce NA\<rbrace> # evs1 \<in> yahalom" | 
| 6400 | 35 | |
| 36 | (*Bob's response to Alice's message.*) | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 37 | | YM2: "\<lbrakk>evs2 \<in> yahalom; Nonce NB \<notin> used evs2; | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 38 | Gets B \<lbrace>Agent A, Nonce NA\<rbrace> \<in> set evs2\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 39 | \<Longrightarrow> Says B Server | 
| 61956 | 40 | \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace> | 
| 11251 | 41 | # evs2 \<in> yahalom" | 
| 6400 | 42 | |
| 43 | (*The Server receives Bob's message. He responds by sending a | |
| 44 | new session key to Alice, with a packet for forwarding to Bob.*) | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 45 | | YM3: "\<lbrakk>evs3 \<in> yahalom; Key KAB \<notin> used evs3; KAB \<in> symKeys; | 
| 11251 | 46 | Gets Server | 
| 61956 | 47 | \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace> | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 48 | \<in> set evs3\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 49 | \<Longrightarrow> Says Server A | 
| 61956 | 50 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key KAB, Nonce NA, Nonce NB\<rbrace>, | 
| 51 | Crypt (shrK B) \<lbrace>Agent A, Key KAB\<rbrace>\<rbrace> | |
| 11251 | 52 | # evs3 \<in> yahalom" | 
| 6400 | 53 | |
| 54 | (*Alice receives the Server's (?) message, checks her Nonce, and | |
| 55 | uses the new session key to send Bob his Nonce. The premise | |
| 11251 | 56 | A \<noteq> Server is needed to prove Says_Server_not_range.*) | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 57 | | YM4: "\<lbrakk>evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys; | 
| 61956 | 58 | Gets A \<lbrace>Crypt(shrK A) \<lbrace>Agent B, Key K, Nonce NA, Nonce NB\<rbrace>, X\<rbrace> | 
| 11251 | 59 | \<in> set evs4; | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 60 | Says A B \<lbrace>Agent A, Nonce NA\<rbrace> \<in> set evs4\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 61 | \<Longrightarrow> Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> # evs4 \<in> yahalom" | 
| 11251 | 62 | |
| 63 | ||
| 64 | declare Says_imp_knows_Spy [THEN analz.Inj, dest] | |
| 65 | declare parts.Body [dest] | |
| 66 | declare Fake_parts_insert_in_Un [dest] | |
| 67 | declare analz_into_parts [dest] | |
| 68 | ||
| 69 | ||
| 61830 | 70 | text\<open>A "possibility property": there are traces that reach the end\<close> | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 71 | lemma "\<lbrakk>A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 72 | \<Longrightarrow> \<exists>X NB. \<exists>evs \<in> yahalom. | 
| 61956 | 73 | Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> \<in> set evs" | 
| 11251 | 74 | apply (intro exI bexI) | 
| 75 | apply (rule_tac [2] yahalom.Nil | |
| 76 | [THEN yahalom.YM1, THEN yahalom.Reception, | |
| 77 | THEN yahalom.YM2, THEN yahalom.Reception, | |
| 78 | THEN yahalom.YM3, THEN yahalom.Reception, | |
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
13956diff
changeset | 79 | THEN yahalom.YM4]) | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
13956diff
changeset | 80 | apply (possibility, simp add: used_Cons) | 
| 11251 | 81 | done | 
| 82 | ||
| 61830 | 83 | subsection\<open>Regularity Lemmas for Yahalom\<close> | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 84 | |
| 11251 | 85 | lemma Gets_imp_Says: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 86 | "\<lbrakk>Gets B X \<in> set evs; evs \<in> yahalom\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs" | 
| 11251 | 87 | by (erule rev_mp, erule yahalom.induct, auto) | 
| 88 | ||
| 89 | (*Must be proved separately for each protocol*) | |
| 90 | lemma Gets_imp_knows_Spy: | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 91 | "\<lbrakk>Gets B X \<in> set evs; evs \<in> yahalom\<rbrakk> \<Longrightarrow> X \<in> knows Spy evs" | 
| 11251 | 92 | by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) | 
| 93 | ||
| 94 | declare Gets_imp_knows_Spy [THEN analz.Inj, dest] | |
| 95 | ||
| 96 | ||
| 61830 | 97 | subsection\<open>For reasoning about the encrypted portion of messages\<close> | 
| 11251 | 98 | |
| 61830 | 99 | text\<open>Lets us treat YM4 using a similar argument as for the Fake case.\<close> | 
| 11251 | 100 | lemma YM4_analz_knows_Spy: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 101 | "\<lbrakk>Gets A \<lbrace>Crypt (shrK A) Y, X\<rbrace> \<in> set evs; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 102 | \<Longrightarrow> X \<in> analz (knows Spy evs)" | 
| 11251 | 103 | by blast | 
| 104 | ||
| 105 | lemmas YM4_parts_knows_Spy = | |
| 45605 | 106 | YM4_analz_knows_Spy [THEN analz_into_parts] | 
| 11251 | 107 | |
| 108 | ||
| 69597 | 109 | text\<open>Theorems of the form \<^term>\<open>X \<notin> parts (knows Spy evs)\<close> imply | 
| 61830 | 110 | that NOBODY sends messages containing X!\<close> | 
| 11251 | 111 | |
| 61830 | 112 | text\<open>Spy never sees a good agent's shared key!\<close> | 
| 11251 | 113 | lemma Spy_see_shrK [simp]: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 114 | "evs \<in> yahalom \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)" | 
| 11251 | 115 | apply (erule yahalom.induct, force, | 
| 13507 | 116 | drule_tac [6] YM4_parts_knows_Spy, simp_all, blast+) | 
| 11251 | 117 | done | 
| 118 | ||
| 119 | lemma Spy_analz_shrK [simp]: | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 120 | "evs \<in> yahalom \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)" | 
| 11251 | 121 | by auto | 
| 122 | ||
| 123 | lemma Spy_see_shrK_D [dest!]: | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 124 | "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs); evs \<in> yahalom\<rbrakk> \<Longrightarrow> A \<in> bad" | 
| 11251 | 125 | by (blast dest: Spy_see_shrK) | 
| 126 | ||
| 61830 | 127 | text\<open>Nobody can have used non-existent keys! | 
| 128 | Needed to apply \<open>analz_insert_Key\<close>\<close> | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 129 | lemma new_keys_not_used [simp]: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 130 | "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 131 | \<Longrightarrow> K \<notin> keysFor (parts (spies evs))" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 132 | apply (erule rev_mp) | 
| 11251 | 133 | apply (erule yahalom.induct, force, | 
| 134 | frule_tac [6] YM4_parts_knows_Spy, simp_all) | |
| 61830 | 135 | txt\<open>Fake\<close> | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 136 | apply (force dest!: keysFor_parts_insert, auto) | 
| 11251 | 137 | done | 
| 138 | ||
| 139 | ||
| 61830 | 140 | subsection\<open>Secrecy Theorems\<close> | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 141 | |
| 11251 | 142 | (**** | 
| 143 | The following is to prove theorems of the form | |
| 144 | ||
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 145 | Key K \<in> analz (insert (Key KAB) (knows Spy evs)) \<Longrightarrow> | 
| 11251 | 146 | Key K \<in> analz (knows Spy evs) | 
| 147 | ||
| 148 | A more general formula must be proved inductively. | |
| 149 | ****) | |
| 150 | ||
| 61830 | 151 | subsection\<open>Session keys are not used to encrypt other session keys\<close> | 
| 11251 | 152 | |
| 153 | lemma analz_image_freshK [rule_format]: | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 154 | "evs \<in> yahalom \<Longrightarrow> | 
| 67613 | 155 | \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow> | 
| 156 | (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) = | |
| 11251 | 157 | (K \<in> KK | Key K \<in> analz (knows Spy evs))" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 158 | by (erule yahalom.induct, | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 159 | drule_tac [7] YM4_analz_knows_Spy, analz_freshK, spy_analz, blast) | 
| 11251 | 160 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 161 | lemma analz_insert_freshK: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 162 | "\<lbrakk>evs \<in> yahalom; KAB \<notin> range shrK\<rbrakk> \<Longrightarrow> | 
| 11655 | 163 | (Key K \<in> analz (insert (Key KAB) (knows Spy evs))) = | 
| 11251 | 164 | (K = KAB | Key K \<in> analz (knows Spy evs))" | 
| 165 | by (simp only: analz_image_freshK analz_image_freshK_simps) | |
| 166 | ||
| 167 | ||
| 61830 | 168 | text\<open>The Key K uniquely identifies the Server's message.\<close> | 
| 11251 | 169 | lemma unique_session_keys: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 170 | "\<lbrakk>Says Server A | 
| 61956 | 171 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>, X\<rbrace> \<in> set evs; | 
| 11251 | 172 | Says Server A' | 
| 61956 | 173 | \<lbrace>Crypt (shrK A') \<lbrace>Agent B', Key K, na', nb'\<rbrace>, X'\<rbrace> \<in> set evs; | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 174 | evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 175 | \<Longrightarrow> A=A' \<and> B=B' \<and> na=na' \<and> nb=nb'" | 
| 11251 | 176 | apply (erule rev_mp, erule rev_mp) | 
| 177 | apply (erule yahalom.induct, simp_all) | |
| 61830 | 178 | txt\<open>YM3, by freshness, and YM4\<close> | 
| 11251 | 179 | apply blast+ | 
| 180 | done | |
| 181 | ||
| 182 | ||
| 61830 | 183 | text\<open>Crucial secrecy property: Spy does not see the keys sent in msg YM3\<close> | 
| 11251 | 184 | lemma secrecy_lemma: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 185 | "\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 186 | \<Longrightarrow> Says Server A | 
| 61956 | 187 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>, | 
| 188 | Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace> | |
| 67613 | 189 | \<in> set evs \<longrightarrow> | 
| 11251 | 190 | Key K \<notin> analz (knows Spy evs)" | 
| 191 | apply (erule yahalom.induct, force, drule_tac [6] YM4_analz_knows_Spy) | |
| 13507 | 192 | apply (simp_all add: pushes analz_insert_eq analz_insert_freshK, spy_analz) (*Fake*) | 
| 11251 | 193 | apply (blast dest: unique_session_keys) (*YM3*) | 
| 194 | done | |
| 195 | ||
| 61830 | 196 | text\<open>Final version\<close> | 
| 11251 | 197 | lemma Spy_not_see_encrypted_key: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 198 | "\<lbrakk>Says Server A | 
| 61956 | 199 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>, | 
| 200 | Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace> | |
| 11251 | 201 | \<in> set evs; | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 202 | A \<notin> bad; B \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 203 | \<Longrightarrow> Key K \<notin> analz (knows Spy evs)" | 
| 11251 | 204 | by (blast dest: secrecy_lemma) | 
| 205 | ||
| 6400 | 206 | |
| 61830 | 207 | subsection\<open>Security Guarantee for A upon receiving YM3\<close> | 
| 11251 | 208 | |
| 61830 | 209 | text\<open>If the encrypted message appears then it originated with the Server\<close> | 
| 11251 | 210 | lemma A_trusts_YM3: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 211 | "\<lbrakk>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace> \<in> parts (knows Spy evs); | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 212 | A \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 213 | \<Longrightarrow> Says Server A | 
| 61956 | 214 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>, | 
| 215 | Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace> | |
| 11251 | 216 | \<in> set evs" | 
| 217 | apply (erule rev_mp) | |
| 218 | apply (erule yahalom.induct, force, | |
| 219 | frule_tac [6] YM4_parts_knows_Spy, simp_all) | |
| 61830 | 220 | txt\<open>Fake, YM3\<close> | 
| 11251 | 221 | apply blast+ | 
| 222 | done | |
| 223 | ||
| 61830 | 224 | text\<open>The obvious combination of \<open>A_trusts_YM3\<close> with | 
| 225 | \<open>Spy_not_see_encrypted_key\<close>\<close> | |
| 11251 | 226 | lemma A_gets_good_key: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 227 | "\<lbrakk>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace> \<in> parts (knows Spy evs); | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 228 | A \<notin> bad; B \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 229 | \<Longrightarrow> Key K \<notin> analz (knows Spy evs)" | 
| 11251 | 230 | by (blast dest!: A_trusts_YM3 Spy_not_see_encrypted_key) | 
| 231 | ||
| 61830 | 232 | subsection\<open>Security Guarantees for B upon receiving YM4\<close> | 
| 11251 | 233 | |
| 61830 | 234 | text\<open>B knows, by the first part of A's message, that the Server distributed | 
| 235 | the key for A and B. But this part says nothing about nonces.\<close> | |
| 11251 | 236 | lemma B_trusts_YM4_shrK: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 237 | "\<lbrakk>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace> \<in> parts (knows Spy evs); | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 238 | B \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 239 | \<Longrightarrow> \<exists>NA NB. Says Server A | 
| 61956 | 240 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, Nonce NA, Nonce NB\<rbrace>, | 
| 241 | Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace> | |
| 11251 | 242 | \<in> set evs" | 
| 243 | apply (erule rev_mp) | |
| 244 | apply (erule yahalom.induct, force, | |
| 245 | frule_tac [6] YM4_parts_knows_Spy, simp_all) | |
| 61830 | 246 | txt\<open>Fake, YM3\<close> | 
| 11251 | 247 | apply blast+ | 
| 248 | done | |
| 249 | ||
| 61830 | 250 | subsection\<open>The Flaw in the Model\<close> | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 251 | |
| 61830 | 252 | text\<open>Up to now, the reasoning is similar to standard Yahalom. Now the | 
| 11251 | 253 | doubtful reasoning occurs. We should not be assuming that an unknown | 
| 254 | key is secure, but the model allows us to: there is no Oops rule to | |
| 61830 | 255 | let session keys become compromised.\<close> | 
| 11251 | 256 | |
| 61830 | 257 | text\<open>B knows, by the second part of A's message, that the Server distributed | 
| 11251 | 258 | the key quoting nonce NB. This part says nothing about agent names. | 
| 259 | Secrecy of K is assumed; the valid Yahalom proof uses (and later proves) | |
| 61830 | 260 | the secrecy of NB.\<close> | 
| 11251 | 261 | lemma B_trusts_YM4_newK [rule_format]: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 262 | "\<lbrakk>Key K \<notin> analz (knows Spy evs); evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 263 | \<Longrightarrow> Crypt K (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow> | 
| 11251 | 264 | (\<exists>A B NA. Says Server A | 
| 61956 | 265 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, | 
| 266 | Nonce NA, Nonce NB\<rbrace>, | |
| 267 | Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace> | |
| 11251 | 268 | \<in> set evs)" | 
| 269 | apply (erule rev_mp) | |
| 270 | apply (erule yahalom.induct, force, | |
| 271 | frule_tac [6] YM4_parts_knows_Spy) | |
| 272 | apply (analz_mono_contra, simp_all) | |
| 61830 | 273 | txt\<open>Fake\<close> | 
| 11251 | 274 | apply blast | 
| 61830 | 275 | txt\<open>YM3\<close> | 
| 11251 | 276 | apply blast | 
| 61830 | 277 | txt\<open>A is uncompromised because NB is secure | 
| 278 | A's certificate guarantees the existence of the Server message\<close> | |
| 11251 | 279 | apply (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad | 
| 280 | dest: Says_imp_spies | |
| 281 | parts.Inj [THEN parts.Fst, THEN A_trusts_YM3]) | |
| 282 | done | |
| 283 | ||
| 284 | ||
| 61830 | 285 | text\<open>B's session key guarantee from YM4. The two certificates contribute to a | 
| 286 | single conclusion about the Server's message.\<close> | |
| 11251 | 287 | lemma B_trusts_YM4: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 288 | "\<lbrakk>Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>, | 
| 61956 | 289 | Crypt K (Nonce NB)\<rbrace> \<in> set evs; | 
| 11251 | 290 | Says B Server | 
| 61956 | 291 | \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace> | 
| 11251 | 292 | \<in> set evs; | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 293 | A \<notin> bad; B \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 294 | \<Longrightarrow> \<exists>na nb. Says Server A | 
| 61956 | 295 | \<lbrace>Crypt (shrK A) \<lbrace>Agent B, Key K, na, nb\<rbrace>, | 
| 296 | Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>\<rbrace> | |
| 11251 | 297 | \<in> set evs" | 
| 298 | by (blast dest: B_trusts_YM4_newK B_trusts_YM4_shrK Spy_not_see_encrypted_key | |
| 299 | unique_session_keys) | |
| 300 | ||
| 301 | ||
| 61830 | 302 | text\<open>The obvious combination of \<open>B_trusts_YM4\<close> with | 
| 303 | \<open>Spy_not_see_encrypted_key\<close>\<close> | |
| 11251 | 304 | lemma B_gets_good_key: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 305 | "\<lbrakk>Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>, | 
| 61956 | 306 | Crypt K (Nonce NB)\<rbrace> \<in> set evs; | 
| 11251 | 307 | Says B Server | 
| 61956 | 308 | \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace> | 
| 11251 | 309 | \<in> set evs; | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 310 | A \<notin> bad; B \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 311 | \<Longrightarrow> Key K \<notin> analz (knows Spy evs)" | 
| 11251 | 312 | by (blast dest!: B_trusts_YM4 Spy_not_see_encrypted_key) | 
| 313 | ||
| 314 | ||
| 315 | (*** Authenticating B to A: these proofs are not considered. | |
| 316 | They are irrelevant to showing the need for Oops. ***) | |
| 317 | ||
| 318 | ||
| 319 | (*** Authenticating A to B using the certificate Crypt K (Nonce NB) ***) | |
| 320 | ||
| 61830 | 321 | text\<open>Assuming the session key is secure, if both certificates are present then | 
| 11251 | 322 | A has said NB. We can't be sure about the rest of A's message, but only | 
| 61830 | 323 | NB matters for freshness.\<close> | 
| 11251 | 324 | lemma A_Said_YM3_lemma [rule_format]: | 
| 325 | "evs \<in> yahalom | |
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 326 | \<Longrightarrow> Key K \<notin> analz (knows Spy evs) \<longrightarrow> | 
| 67613 | 327 | Crypt K (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow> | 
| 328 | Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow> | |
| 329 | B \<notin> bad \<longrightarrow> | |
| 61956 | 330 | (\<exists>X. Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> \<in> set evs)" | 
| 11251 | 331 | apply (erule yahalom.induct, force, | 
| 332 | frule_tac [6] YM4_parts_knows_Spy) | |
| 333 | apply (analz_mono_contra, simp_all) | |
| 61830 | 334 | txt\<open>Fake\<close> | 
| 11251 | 335 | apply blast | 
| 61830 | 336 | txt\<open>YM3: by \<open>new_keys_not_used\<close>, the message | 
| 69597 | 337 | \<^term>\<open>Crypt K (Nonce NB)\<close> could not exist\<close> | 
| 11251 | 338 | apply (force dest!: Crypt_imp_keysFor) | 
| 69597 | 339 | txt\<open>YM4: was \<^term>\<open>Crypt K (Nonce NB)\<close> the very last message? | 
| 61830 | 340 | If not, use the induction hypothesis\<close> | 
| 11251 | 341 | apply (simp add: ex_disj_distrib) | 
| 61830 | 342 | txt\<open>yes: apply unicity of session keys\<close> | 
| 11251 | 343 | apply (blast dest!: Gets_imp_Says A_trusts_YM3 B_trusts_YM4_shrK | 
| 344 | Crypt_Spy_analz_bad | |
| 345 | dest: Says_imp_knows_Spy [THEN parts.Inj] unique_session_keys) | |
| 346 | done | |
| 347 | ||
| 61830 | 348 | text\<open>If B receives YM4 then A has used nonce NB (and therefore is alive). | 
| 11251 | 349 | Moreover, A associates K with NB (thus is talking about the same run). | 
| 61830 | 350 | Other premises guarantee secrecy of K.\<close> | 
| 11251 | 351 | lemma YM4_imp_A_Said_YM3 [rule_format]: | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 352 | "\<lbrakk>Gets B \<lbrace>Crypt (shrK B) \<lbrace>Agent A, Key K\<rbrace>, | 
| 61956 | 353 | Crypt K (Nonce NB)\<rbrace> \<in> set evs; | 
| 11251 | 354 | Says B Server | 
| 61956 | 355 | \<lbrace>Agent B, Nonce NB, Crypt (shrK B) \<lbrace>Agent A, Nonce NA\<rbrace>\<rbrace> | 
| 11251 | 356 | \<in> set evs; | 
| 76287 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 357 | A \<notin> bad; B \<notin> bad; evs \<in> yahalom\<rbrakk> | 
| 
cdc14f94c754
Elimination of the archaic ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 358 | \<Longrightarrow> \<exists>X. Says A B \<lbrace>X, Crypt K (Nonce NB)\<rbrace> \<in> set evs" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 359 | by (blast intro!: A_Said_YM3_lemma | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 360 | dest: Spy_not_see_encrypted_key B_trusts_YM4 Gets_imp_Says) | 
| 6400 | 361 | |
| 362 | end |