| author | webertj | 
| Wed, 10 Mar 2004 22:33:48 +0100 | |
| changeset 14456 | cca28ec5f9a6 | 
| parent 14218 | db95d1c2f51b | 
| child 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 14199 | 1 | (* Title: HOL/Auth/SET/Cardholder_Registration | 
| 2 | ID: $Id$ | |
| 3 | Authors: Giampaolo Bella, Fabio Massacci, Lawrence C Paulson, | |
| 4 | Piero Tramontano | |
| 5 | *) | |
| 6 | ||
| 7 | header{*The SET Cardholder Registration Protocol*}
 | |
| 8 | ||
| 9 | theory Cardholder_Registration = PublicSET: | |
| 10 | ||
| 11 | text{*Note: nonces seem to consist of 20 bytes.  That includes both freshness
 | |
| 12 | challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret) | |
| 13 | *} | |
| 14 | ||
| 15 | text{*Simplifications involving @{text analz_image_keys_simps} appear to
 | |
| 16 | have become much slower. The cause is unclear. However, there is a big blow-up | |
| 17 | and the rewriting is very sensitive to the set of rewrite rules given.*} | |
| 18 | ||
| 19 | subsection{*Predicate Formalizing the Encryption Association between Keys *}
 | |
| 20 | ||
| 21 | consts | |
| 22 | KeyCryptKey :: "[key, key, event list] => bool" | |
| 23 | ||
| 24 | primrec | |
| 25 | ||
| 26 | KeyCryptKey_Nil: | |
| 27 | "KeyCryptKey DK K [] = False" | |
| 28 | ||
| 29 | KeyCryptKey_Cons: | |
| 30 |       --{*Says is the only important case.
 | |
| 31 | 1st case: CR5, where KC3 encrypts KC2. | |
| 32 | 2nd case: any use of priEK C. | |
| 33 | Revision 1.12 has a more complicated version with separate treatment of | |
| 34 | the dependency of KC1, KC2 and KC3 on priEK (CA i.) Not needed since | |
| 35 | priEK C is never sent (and so can't be lost except at the start). *} | |
| 36 | "KeyCryptKey DK K (ev # evs) = | |
| 37 | (KeyCryptKey DK K evs | | |
| 38 | (case ev of | |
| 39 | Says A B Z => | |
| 40 | ((\<exists>N X Y. A \<noteq> Spy & | |
| 41 | DK \<in> symKeys & | |
| 42 | 		 Z = {|Crypt DK {|Agent A, Nonce N, Key K, X|}, Y|}) |
 | |
| 43 | (\<exists>C. DK = priEK C)) | |
| 44 | | Gets A' X => False | |
| 45 | | Notes A' X => False))" | |
| 46 | ||
| 47 | ||
| 48 | subsection{*Predicate formalizing the association between keys and nonces *}
 | |
| 49 | ||
| 50 | consts | |
| 51 | KeyCryptNonce :: "[key, key, event list] => bool" | |
| 52 | ||
| 53 | primrec | |
| 54 | ||
| 55 | KeyCryptNonce_Nil: | |
| 56 | "KeyCryptNonce EK K [] = False" | |
| 57 | ||
| 58 | KeyCryptNonce_Cons: | |
| 59 |   --{*Says is the only important case.
 | |
| 60 | 1st case: CR3, where KC1 encrypts NC2 (distinct from CR5 due to EXH); | |
| 61 | 2nd case: CR5, where KC3 encrypts NC3; | |
| 62 | 3rd case: CR6, where KC2 encrypts NC3; | |
| 63 | 4th case: CR6, where KC2 encrypts NonceCCA; | |
| 64 |     5th case: any use of @{term "priEK C"} (including CardSecret).
 | |
| 65 | NB the only Nonces we need to keep secret are CardSecret and NonceCCA. | |
| 66 |     But we can't prove @{text Nonce_compromise} unless the relation covers ALL
 | |
| 67 | nonces that the protocol keeps secret. | |
| 68 | *} | |
| 69 | "KeyCryptNonce DK N (ev # evs) = | |
| 70 | (KeyCryptNonce DK N evs | | |
| 71 | (case ev of | |
| 72 | Says A B Z => | |
| 73 | A \<noteq> Spy & | |
| 74 | ((\<exists>X Y. DK \<in> symKeys & | |
| 75 | 	       Z = (EXHcrypt DK X {|Agent A, Nonce N|} Y)) |
 | |
| 76 | (\<exists>X Y. DK \<in> symKeys & | |
| 77 | 	       Z = {|Crypt DK {|Agent A, Nonce N, X|}, Y|}) |
 | |
| 78 | (\<exists>K i X Y. | |
| 79 | K \<in> symKeys & | |
| 80 |           Z = Crypt K {|sign (priSK (CA i)) {|Agent B, Nonce N, X|}, Y|} &
 | |
| 81 | (DK=K | KeyCryptKey DK K evs)) | | |
| 82 | (\<exists>K C NC3 Y. | |
| 83 | K \<in> symKeys & | |
| 84 | Z = Crypt K | |
| 85 |  	        {|sign (priSK C) {|Agent B, Nonce NC3, Agent C, Nonce N|},
 | |
| 86 | Y|} & | |
| 87 | (DK=K | KeyCryptKey DK K evs)) | | |
| 88 | (\<exists>C. DK = priEK C)) | |
| 89 | | Gets A' X => False | |
| 90 | | Notes A' X => False))" | |
| 91 | ||
| 92 | ||
| 93 | subsection{*Formal protocol definition *}
 | |
| 94 | ||
| 95 | consts set_cr :: "event list set" | |
| 96 | inductive set_cr | |
| 97 | intros | |
| 98 | ||
| 99 |   Nil:    --{*Initial trace is empty*}
 | |
| 100 | "[] \<in> set_cr" | |
| 101 | ||
| 102 |   Fake:    --{*The spy MAY say anything he CAN say.*}
 | |
| 103 | "[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |] | |
| 104 | ==> Says Spy B X # evsf \<in> set_cr" | |
| 105 | ||
| 106 |   Reception: --{*If A sends a message X to B, then B might receive it*}
 | |
| 107 | "[| evsr \<in> set_cr; Says A B X \<in> set evsr |] | |
| 108 | ==> Gets B X # evsr \<in> set_cr" | |
| 109 | ||
| 110 |   SET_CR1: --{*CardCInitReq: C initiates a run, sending a nonce to CCA*}
 | |
| 111 | "[| evs1 \<in> set_cr; C = Cardholder k; Nonce NC1 \<notin> used evs1 |] | |
| 112 | 	      ==> Says C (CA i) {|Agent C, Nonce NC1|} # evs1 \<in> set_cr"
 | |
| 113 | ||
| 114 |   SET_CR2: --{*CardCInitRes: CA responds sending NC1 and its certificates*}
 | |
| 115 | "[| evs2 \<in> set_cr; | |
| 116 | 		 Gets (CA i) {|Agent C, Nonce NC1|} \<in> set evs2 |]
 | |
| 117 | ==> Says (CA i) C | |
| 118 | 		       {|sign (priSK (CA i)) {|Agent C, Nonce NC1|},
 | |
| 119 | cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA), | |
| 120 | cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} | |
| 121 | # evs2 \<in> set_cr" | |
| 122 | ||
| 123 | SET_CR3: | |
| 124 |    --{*RegFormReq: C sends his PAN and a new nonce to CA.
 | |
| 125 | C verifies that | |
| 126 | - nonce received is the same as that sent; | |
| 127 | - certificates are signed by RCA; | |
| 128 | - certificates are an encryption certificate (flag is onlyEnc) and a | |
| 129 | signature certificate (flag is onlySig); | |
| 130 | - certificates pertain to the CA that C contacted (this is done by | |
| 131 | checking the signature). | |
| 132 | C generates a fresh symmetric key KC1. | |
| 133 |    The point of encrypting @{term "{|Agent C, Nonce NC2, Hash (Pan(pan C))|}"}
 | |
| 134 | is not clear. *} | |
| 135 | "[| evs3 \<in> set_cr; C = Cardholder k; | |
| 136 | Nonce NC2 \<notin> used evs3; | |
| 137 | Key KC1 \<notin> used evs3; KC1 \<in> symKeys; | |
| 138 |     Gets C {|sign (invKey SKi) {|Agent X, Nonce NC1|},
 | |
| 139 | cert (CA i) EKi onlyEnc (priSK RCA), | |
| 140 | cert (CA i) SKi onlySig (priSK RCA)|} | |
| 141 | \<in> set evs3; | |
| 142 |     Says C (CA i) {|Agent C, Nonce NC1|} \<in> set evs3|]
 | |
| 143 |  ==> Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
 | |
| 144 |        # Notes C {|Key KC1, Agent (CA i)|}
 | |
| 145 | # evs3 \<in> set_cr" | |
| 146 | ||
| 147 | SET_CR4: | |
| 148 |     --{*RegFormRes:
 | |
| 149 | CA responds sending NC2 back with a new nonce NCA, after checking that | |
| 150 |      - the digital envelope is correctly encrypted by @{term "pubEK (CA i)"}
 | |
| 151 | - the entire message is encrypted with the same key found inside the | |
| 152 | envelope (here, KC1) *} | |
| 153 | "[| evs4 \<in> set_cr; | |
| 154 | Nonce NCA \<notin> used evs4; KC1 \<in> symKeys; | |
| 155 |     Gets (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan X)))
 | |
| 156 | \<in> set evs4 |] | |
| 157 | ==> Says (CA i) C | |
| 158 | 	  {|sign (priSK (CA i)) {|Agent C, Nonce NC2, Nonce NCA|},
 | |
| 159 | cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA), | |
| 160 | cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} | |
| 161 | # evs4 \<in> set_cr" | |
| 162 | ||
| 163 | SET_CR5: | |
| 164 |    --{*CertReq: C sends his PAN, a new nonce, its proposed public signature key
 | |
| 165 | and its half of the secret value to CA. | |
| 166 | We now assume that C has a fixed key pair, and he submits (pubSK C). | |
| 167 | The protocol does not require this key to be fresh. | |
| 168 | The encryption below is actually EncX.*} | |
| 169 | "[| evs5 \<in> set_cr; C = Cardholder k; | |
| 170 | Nonce NC3 \<notin> used evs5; Nonce CardSecret \<notin> used evs5; NC3\<noteq>CardSecret; | |
| 171 | Key KC2 \<notin> used evs5; KC2 \<in> symKeys; | |
| 172 | Key KC3 \<notin> used evs5; KC3 \<in> symKeys; KC2\<noteq>KC3; | |
| 173 |     Gets C {|sign (invKey SKi) {|Agent C, Nonce NC2, Nonce NCA|},
 | |
| 174 | cert (CA i) EKi onlyEnc (priSK RCA), | |
| 175 | cert (CA i) SKi onlySig (priSK RCA) |} | |
| 176 | \<in> set evs5; | |
| 177 |     Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
 | |
| 178 | \<in> set evs5 |] | |
| 179 | ==> Says C (CA i) | |
| 180 |          {|Crypt KC3
 | |
| 181 | 	     {|Agent C, Nonce NC3, Key KC2, Key (pubSK C),
 | |
| 182 | Crypt (priSK C) | |
| 183 | 	         (Hash {|Agent C, Nonce NC3, Key KC2,
 | |
| 184 | Key (pubSK C), Pan (pan C), Nonce CardSecret|})|}, | |
| 185 |            Crypt EKi {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
 | |
| 186 |     # Notes C {|Key KC2, Agent (CA i)|}
 | |
| 187 |     # Notes C {|Key KC3, Agent (CA i)|}
 | |
| 188 | # evs5 \<in> set_cr" | |
| 189 | ||
| 190 | ||
| 191 |   --{* CertRes: CA responds sending NC3 back with its half of the secret value,
 | |
| 192 | its signature certificate and the new cardholder signature | |
| 193 | certificate. CA checks to have never certified the key proposed by C. | |
| 194 | NOTE: In Merchant Registration, the corresponding rule (4) | |
| 195 |    uses the "sign" primitive. The encryption below is actually @{term EncK}, 
 | |
| 196 |    which is just @{term "Crypt K (sign SK X)"}.
 | |
| 197 | *} | |
| 198 | ||
| 199 | SET_CR6: | |
| 200 | "[| evs6 \<in> set_cr; | |
| 201 | Nonce NonceCCA \<notin> used evs6; | |
| 202 | KC2 \<in> symKeys; KC3 \<in> symKeys; cardSK \<notin> symKeys; | |
| 203 | Notes (CA i) (Key cardSK) \<notin> set evs6; | |
| 204 | Gets (CA i) | |
| 205 |       {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, Key cardSK,
 | |
| 206 | Crypt (invKey cardSK) | |
| 207 |                       (Hash {|Agent C, Nonce NC3, Key KC2,
 | |
| 208 | Key cardSK, Pan (pan C), Nonce CardSecret|})|}, | |
| 209 |         Crypt (pubEK (CA i)) {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
 | |
| 210 | \<in> set evs6 |] | |
| 211 | ==> Says (CA i) C | |
| 212 | (Crypt KC2 | |
| 213 | 	  {|sign (priSK (CA i))
 | |
| 214 | 	         {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
 | |
| 215 | certC (pan C) cardSK (XOR(CardSecret,NonceCCA)) onlySig (priSK (CA i)), | |
| 216 | cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}) | |
| 217 | # Notes (CA i) (Key cardSK) | |
| 218 | # evs6 \<in> set_cr" | |
| 219 | ||
| 220 | ||
| 221 | declare Says_imp_knows_Spy [THEN parts.Inj, dest] | |
| 222 | declare parts.Body [dest] | |
| 223 | declare analz_into_parts [dest] | |
| 224 | declare Fake_parts_insert_in_Un [dest] | |
| 225 | ||
| 226 | text{*A "possibility property": there are traces that reach the end.
 | |
| 227 | An unconstrained proof with many subgoals.*} | |
| 228 | ||
| 229 | lemma Says_to_Gets: | |
| 230 | "Says A B X # evs \<in> set_cr ==> Gets B X # Says A B X # evs \<in> set_cr" | |
| 231 | by (rule set_cr.Reception, auto) | |
| 232 | ||
| 233 | text{*The many nonces and keys generated, some simultaneously, force us to
 | |
| 234 | introduce them explicitly as shown below.*} | |
| 235 | lemma possibility_CR6: | |
| 236 | "[|NC1 < (NC2::nat); NC2 < NC3; NC3 < NCA ; | |
| 237 | NCA < NonceCCA; NonceCCA < CardSecret; | |
| 238 | KC1 < (KC2::key); KC2 < KC3; | |
| 239 | KC1 \<in> symKeys; Key KC1 \<notin> used []; | |
| 240 | KC2 \<in> symKeys; Key KC2 \<notin> used []; | |
| 241 | KC3 \<in> symKeys; Key KC3 \<notin> used []; | |
| 242 | C = Cardholder k|] | |
| 243 | ==> \<exists>evs \<in> set_cr. | |
| 244 | Says (CA i) C | |
| 245 | (Crypt KC2 | |
| 246 |              {|sign (priSK (CA i))
 | |
| 247 |                     {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
 | |
| 248 | certC (pan C) (pubSK (Cardholder k)) (XOR(CardSecret,NonceCCA)) | |
| 249 | onlySig (priSK (CA i)), | |
| 250 | cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}) | |
| 251 | \<in> set evs" | |
| 252 | apply (intro exI bexI) | |
| 14206 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 253 | apply (rule_tac [2] | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 254 | set_cr.Nil | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 255 | [THEN set_cr.SET_CR1 [of concl: C i NC1], | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 256 | THEN Says_to_Gets, | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 257 | THEN set_cr.SET_CR2 [of concl: i C NC1], | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 258 | THEN Says_to_Gets, | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 259 | THEN set_cr.SET_CR3 [of concl: C i KC1 _ NC2], | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 260 | THEN Says_to_Gets, | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 261 | THEN set_cr.SET_CR4 [of concl: i C NC2 NCA], | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 262 | THEN Says_to_Gets, | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 263 | THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret], | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 264 | THEN Says_to_Gets, | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 265 | THEN set_cr.SET_CR6 [of concl: i C KC2]]) | 
| 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 266 | apply (tactic "basic_possibility_tac") | 
| 14199 | 267 | apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq) | 
| 268 | done | |
| 269 | ||
| 270 | text{*General facts about message reception*}
 | |
| 271 | lemma Gets_imp_Says: | |
| 272 | "[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> \<exists>A. Says A B X \<in> set evs" | |
| 273 | apply (erule rev_mp) | |
| 274 | apply (erule set_cr.induct, auto) | |
| 275 | done | |
| 276 | ||
| 277 | lemma Gets_imp_knows_Spy: | |
| 278 | "[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> X \<in> knows Spy evs" | |
| 279 | by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) | |
| 280 | declare Gets_imp_knows_Spy [THEN parts.Inj, dest] | |
| 281 | ||
| 282 | ||
| 283 | subsection{*Proofs on keys *}
 | |
| 284 | ||
| 285 | text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
 | |
| 286 | ||
| 287 | lemma Spy_see_private_Key [simp]: | |
| 288 | "evs \<in> set_cr | |
| 289 | ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)" | |
| 290 | by (erule set_cr.induct, auto) | |
| 291 | ||
| 292 | lemma Spy_analz_private_Key [simp]: | |
| 293 | "evs \<in> set_cr ==> | |
| 294 | (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)" | |
| 295 | by auto | |
| 296 | ||
| 297 | declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!] | |
| 298 | declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!] | |
| 299 | ||
| 300 | ||
| 301 | subsection{*Begin Piero's Theorems on Certificates*}
 | |
| 302 | text{*Trivial in the current model, where certificates by RCA are secure *}
 | |
| 303 | ||
| 304 | lemma Crypt_valid_pubEK: | |
| 305 |      "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|}
 | |
| 306 | \<in> parts (knows Spy evs); | |
| 307 | evs \<in> set_cr |] ==> EKi = pubEK C" | |
| 308 | apply (erule rev_mp) | |
| 309 | apply (erule set_cr.induct, auto) | |
| 310 | done | |
| 311 | ||
| 312 | lemma certificate_valid_pubEK: | |
| 313 | "[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs); | |
| 314 | evs \<in> set_cr |] | |
| 315 | ==> EKi = pubEK C" | |
| 316 | apply (unfold cert_def signCert_def) | |
| 317 | apply (blast dest!: Crypt_valid_pubEK) | |
| 318 | done | |
| 319 | ||
| 320 | lemma Crypt_valid_pubSK: | |
| 321 |      "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|}
 | |
| 322 | \<in> parts (knows Spy evs); | |
| 323 | evs \<in> set_cr |] ==> SKi = pubSK C" | |
| 324 | apply (erule rev_mp) | |
| 325 | apply (erule set_cr.induct, auto) | |
| 326 | done | |
| 327 | ||
| 328 | lemma certificate_valid_pubSK: | |
| 329 | "[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs); | |
| 330 | evs \<in> set_cr |] ==> SKi = pubSK C" | |
| 331 | apply (unfold cert_def signCert_def) | |
| 332 | apply (blast dest!: Crypt_valid_pubSK) | |
| 333 | done | |
| 334 | ||
| 335 | lemma Gets_certificate_valid: | |
| 336 |      "[| Gets A {| X, cert C EKi onlyEnc (priSK RCA),
 | |
| 337 | cert C SKi onlySig (priSK RCA)|} \<in> set evs; | |
| 338 | evs \<in> set_cr |] | |
| 339 | ==> EKi = pubEK C & SKi = pubSK C" | |
| 340 | by (blast dest: certificate_valid_pubEK certificate_valid_pubSK) | |
| 341 | ||
| 342 | text{*Nobody can have used non-existent keys!*}
 | |
| 343 | lemma new_keys_not_used: | |
| 344 | "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr|] | |
| 345 | ==> K \<notin> keysFor (parts (knows Spy evs))" | |
| 346 | apply (erule rev_mp) | |
| 347 | apply (erule rev_mp) | |
| 348 | apply (erule set_cr.induct) | |
| 349 | apply (frule_tac [8] Gets_certificate_valid) | |
| 350 | apply (frule_tac [6] Gets_certificate_valid, simp_all) | |
| 14218 | 351 | apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
 | 
| 352 | apply (blast,auto)  --{*Others*}
 | |
| 14199 | 353 | done | 
| 354 | ||
| 355 | ||
| 356 | subsection{*New versions: as above, but generalized to have the KK argument *}
 | |
| 357 | ||
| 358 | lemma gen_new_keys_not_used: | |
| 359 | "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |] | |
| 360 | ==> Key K \<notin> used evs --> K \<in> symKeys --> | |
| 361 | K \<notin> keysFor (parts (Key`KK Un knows Spy evs))" | |
| 362 | by (auto simp add: new_keys_not_used) | |
| 363 | ||
| 364 | lemma gen_new_keys_not_analzd: | |
| 365 | "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |] | |
| 366 | ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))" | |
| 367 | by (blast intro: keysFor_mono [THEN [2] rev_subsetD] | |
| 368 | dest: gen_new_keys_not_used) | |
| 369 | ||
| 370 | lemma analz_Key_image_insert_eq: | |
| 371 | "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr |] | |
| 372 | ==> analz (Key ` (insert K KK) \<union> knows Spy evs) = | |
| 373 | insert (Key K) (analz (Key ` KK \<union> knows Spy evs))" | |
| 374 | by (simp add: gen_new_keys_not_analzd) | |
| 375 | ||
| 376 | lemma Crypt_parts_imp_used: | |
| 377 | "[|Crypt K X \<in> parts (knows Spy evs); | |
| 378 | K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs" | |
| 379 | apply (rule ccontr) | |
| 380 | apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor) | |
| 381 | done | |
| 382 | ||
| 383 | lemma Crypt_analz_imp_used: | |
| 384 | "[|Crypt K X \<in> analz (knows Spy evs); | |
| 385 | K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs" | |
| 386 | by (blast intro: Crypt_parts_imp_used) | |
| 387 | ||
| 388 | ||
| 14218 | 389 | (*<*) | 
| 14199 | 390 | subsection{*Messages signed by CA*}
 | 
| 391 | ||
| 392 | text{*Message @{text SET_CR2}: C can check CA's signature if he has received
 | |
| 393 | CA's certificate.*} | |
| 394 | lemma CA_Says_2_lemma: | |
| 395 |      "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC1|})
 | |
| 396 | \<in> parts (knows Spy evs); | |
| 397 | evs \<in> set_cr; (CA i) \<notin> bad |] | |
| 398 |      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
 | |
| 399 | \<in> set evs" | |
| 400 | apply (erule rev_mp) | |
| 401 | apply (erule set_cr.induct, auto) | |
| 402 | done | |
| 403 | ||
| 404 | text{*Ever used?*}
 | |
| 405 | lemma CA_Says_2: | |
| 406 |      "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC1|})
 | |
| 407 | \<in> parts (knows Spy evs); | |
| 408 | cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs); | |
| 409 | evs \<in> set_cr; (CA i) \<notin> bad |] | |
| 410 |       ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
 | |
| 411 | \<in> set evs" | |
| 412 | by (blast dest!: certificate_valid_pubSK intro!: CA_Says_2_lemma) | |
| 413 | ||
| 414 | ||
| 415 | text{*Message @{text SET_CR4}: C can check CA's signature if he has received
 | |
| 416 | CA's certificate.*} | |
| 417 | lemma CA_Says_4_lemma: | |
| 418 |      "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
 | |
| 419 | \<in> parts (knows Spy evs); | |
| 420 | evs \<in> set_cr; (CA i) \<notin> bad |] | |
| 421 |       ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
 | |
| 422 |                      {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
 | |
| 423 | apply (erule rev_mp) | |
| 424 | apply (erule set_cr.induct, auto) | |
| 425 | done | |
| 426 | ||
| 427 | text{*NEVER USED*}
 | |
| 428 | lemma CA_Says_4: | |
| 429 |      "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
 | |
| 430 | \<in> parts (knows Spy evs); | |
| 431 | cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs); | |
| 432 | evs \<in> set_cr; (CA i) \<notin> bad |] | |
| 433 |       ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
 | |
| 434 |                    {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
 | |
| 435 | by (blast dest!: certificate_valid_pubSK intro!: CA_Says_4_lemma) | |
| 436 | ||
| 437 | ||
| 438 | text{*Message @{text SET_CR6}: C can check CA's signature if he has
 | |
| 439 | received CA's certificate.*} | |
| 440 | lemma CA_Says_6_lemma: | |
| 441 | "[| Crypt (priSK (CA i)) | |
| 442 |                (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
 | |
| 443 | \<in> parts (knows Spy evs); | |
| 444 | evs \<in> set_cr; (CA i) \<notin> bad |] | |
| 445 |       ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
 | |
| 446 |       {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
 | |
| 447 | apply (erule rev_mp) | |
| 448 | apply (erule set_cr.induct, auto) | |
| 449 | done | |
| 450 | ||
| 451 | text{*NEVER USED*}
 | |
| 452 | lemma CA_Says_6: | |
| 453 |      "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
 | |
| 454 | \<in> parts (knows Spy evs); | |
| 455 | cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs); | |
| 456 | evs \<in> set_cr; (CA i) \<notin> bad |] | |
| 457 |       ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
 | |
| 458 |                     {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
 | |
| 459 | by (blast dest!: certificate_valid_pubSK intro!: CA_Says_6_lemma) | |
| 14218 | 460 | (*>*) | 
| 14199 | 461 | |
| 462 | ||
| 463 | subsection{*Useful lemmas *}
 | |
| 464 | ||
| 465 | text{*Rewriting rule for private encryption keys.  Analogous rewriting rules
 | |
| 466 | for other keys aren't needed.*} | |
| 467 | ||
| 468 | lemma parts_image_priEK: | |
| 469 | "[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs)); | |
| 470 | evs \<in> set_cr|] ==> priEK C \<in> KK | C \<in> bad" | |
| 471 | by auto | |
| 472 | ||
| 473 | text{*trivial proof because (priEK C) never appears even in (parts evs)*}
 | |
| 474 | lemma analz_image_priEK: | |
| 475 | "evs \<in> set_cr ==> | |
| 476 | (Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) = | |
| 477 | (priEK C \<in> KK | C \<in> bad)" | |
| 478 | by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD]) | |
| 479 | ||
| 480 | ||
| 481 | subsection{*Secrecy of Session Keys *}
 | |
| 482 | ||
| 483 | subsubsection{*Lemmas about the predicate KeyCryptKey *}
 | |
| 484 | ||
| 485 | text{*A fresh DK cannot be associated with any other
 | |
| 486 | (with respect to a given trace). *} | |
| 487 | lemma DK_fresh_not_KeyCryptKey: | |
| 488 | "[| Key DK \<notin> used evs; evs \<in> set_cr |] ==> ~ KeyCryptKey DK K evs" | |
| 489 | apply (erule rev_mp) | |
| 490 | apply (erule set_cr.induct) | |
| 491 | apply (simp_all (no_asm_simp)) | |
| 492 | apply (blast dest: Crypt_analz_imp_used)+ | |
| 493 | done | |
| 494 | ||
| 495 | text{*A fresh K cannot be associated with any other.  The assumption that
 | |
| 496 | DK isn't a private encryption key may be an artifact of the particular | |
| 497 | definition of KeyCryptKey.*} | |
| 498 | lemma K_fresh_not_KeyCryptKey: | |
| 499 | "[|\<forall>C. DK \<noteq> priEK C; Key K \<notin> used evs|] ==> ~ KeyCryptKey DK K evs" | |
| 500 | apply (induct evs) | |
| 501 | apply (auto simp add: parts_insert2 split add: event.split) | |
| 502 | done | |
| 503 | ||
| 504 | ||
| 505 | text{*This holds because if (priEK (CA i)) appears in any traffic then it must
 | |
| 506 |   be known to the Spy, by @{term Spy_see_private_Key}*}
 | |
| 507 | lemma cardSK_neq_priEK: | |
| 508 | "[|Key cardSK \<notin> analz (knows Spy evs); | |
| 509 | Key cardSK : parts (knows Spy evs); | |
| 510 | evs \<in> set_cr|] ==> cardSK \<noteq> priEK C" | |
| 511 | by blast | |
| 512 | ||
| 513 | lemma not_KeyCryptKey_cardSK [rule_format (no_asm)]: | |
| 514 | "[|cardSK \<notin> symKeys; \<forall>C. cardSK \<noteq> priEK C; evs \<in> set_cr|] ==> | |
| 515 | Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptKey cardSK K evs" | |
| 516 | by (erule set_cr.induct, analz_mono_contra, auto) | |
| 517 | ||
| 518 | text{*Lemma for message 5: pubSK C is never used to encrypt Keys.*}
 | |
| 519 | lemma pubSK_not_KeyCryptKey [simp]: "~ KeyCryptKey (pubSK C) K evs" | |
| 520 | apply (induct_tac "evs") | |
| 521 | apply (auto simp add: parts_insert2 split add: event.split) | |
| 522 | done | |
| 523 | ||
| 524 | text{*Lemma for message 6: either cardSK is compromised (when we don't care)
 | |
| 525 | or else cardSK hasn't been used to encrypt K. Previously we treated | |
| 526 | message 5 in the same way, but the current model assumes that rule | |
| 527 |   @{text SET_CR5} is executed only by honest agents.*}
 | |
| 528 | lemma msg6_KeyCryptKey_disj: | |
| 529 |      "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
 | |
| 530 | \<in> set evs; | |
| 531 | cardSK \<notin> symKeys; evs \<in> set_cr|] | |
| 532 | ==> Key cardSK \<in> analz (knows Spy evs) | | |
| 533 | (\<forall>K. ~ KeyCryptKey cardSK K evs)" | |
| 534 | by (blast dest: not_KeyCryptKey_cardSK intro: cardSK_neq_priEK) | |
| 535 | ||
| 536 | text{*As usual: we express the property as a logical equivalence*}
 | |
| 537 | lemma Key_analz_image_Key_lemma: | |
| 538 | "P --> (Key K \<in> analz (Key`KK Un H)) --> (K \<in> KK | Key K \<in> analz H) | |
| 539 | ==> | |
| 540 | P --> (Key K \<in> analz (Key`KK Un H)) = (K \<in> KK | Key K \<in> analz H)" | |
| 541 | by (blast intro: analz_mono [THEN [2] rev_subsetD]) | |
| 542 | ||
| 543 | ML | |
| 544 | {*
 | |
| 545 | val Gets_certificate_valid = thm "Gets_certificate_valid"; | |
| 546 | ||
| 547 | fun valid_certificate_tac i = | |
| 548 | EVERY [ftac Gets_certificate_valid i, | |
| 549 | assume_tac i, | |
| 550 | etac conjE i, REPEAT (hyp_subst_tac i)]; | |
| 551 | *} | |
| 552 | ||
| 553 | text{*The @{text "(no_asm)"} attribute is essential, since it retains
 | |
| 554 | the quantifier and allows the simprule's condition to itself be simplified.*} | |
| 555 | lemma symKey_compromise [rule_format (no_asm)]: | |
| 556 | "evs \<in> set_cr ==> | |
| 557 | (\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. ~ KeyCryptKey K SK evs) --> | |
| 558 | (Key SK \<in> analz (Key`KK Un (knows Spy evs))) = | |
| 559 | (SK \<in> KK | Key SK \<in> analz (knows Spy evs)))" | |
| 560 | apply (erule set_cr.induct) | |
| 561 | apply (rule_tac [!] allI) + | |
| 562 | apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+ | |
| 563 | apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
 | |
| 564 | apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
 | |
| 565 | apply (erule_tac [9] msg6_KeyCryptKey_disj [THEN disjE]) | |
| 566 | apply (simp_all | |
| 567 | del: image_insert image_Un imp_disjL | |
| 568 | add: analz_image_keys_simps analz_knows_absorb | |
| 569 | analz_Key_image_insert_eq notin_image_iff | |
| 570 | K_fresh_not_KeyCryptKey | |
| 571 | DK_fresh_not_KeyCryptKey ball_conj_distrib | |
| 572 | analz_image_priEK disj_simps) | |
| 573 |   --{*46 seconds on a 1.8GHz machine*}
 | |
| 574 | apply spy_analz | |
| 14218 | 575 | apply blast  --{*3*}
 | 
| 576 | apply blast  --{*5*}
 | |
| 14199 | 577 | done | 
| 578 | ||
| 579 | text{*The remaining quantifiers seem to be essential.
 | |
| 580 | NO NEED to assume the cardholder's OK: bad cardholders don't do anything | |
| 581 | wrong!!*} | |
| 582 | lemma symKey_secrecy [rule_format]: | |
| 583 | "[|CA i \<notin> bad; K \<in> symKeys; evs \<in> set_cr|] | |
| 584 | ==> \<forall>X c. Says (Cardholder c) (CA i) X \<in> set evs --> | |
| 585 |                 Key K \<in> parts{X} -->
 | |
| 586 | Cardholder c \<notin> bad --> | |
| 587 | Key K \<notin> analz (knows Spy evs)" | |
| 588 | apply (erule set_cr.induct) | |
| 589 | apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
 | |
| 590 | apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
 | |
| 591 | apply (erule_tac [11] msg6_KeyCryptKey_disj [THEN disjE]) | |
| 592 | apply (simp_all del: image_insert image_Un imp_disjL | |
| 593 | add: symKey_compromise fresh_notin_analz_knows_Spy | |
| 594 | analz_image_keys_simps analz_knows_absorb | |
| 595 | analz_Key_image_insert_eq notin_image_iff | |
| 596 | K_fresh_not_KeyCryptKey | |
| 597 | DK_fresh_not_KeyCryptKey | |
| 598 | analz_image_priEK) | |
| 599 |   --{*13 seconds on a 1.8GHz machine*}
 | |
| 14218 | 600 | apply spy_analz  --{*Fake*}
 | 
| 14199 | 601 | apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used) | 
| 602 | done | |
| 603 | ||
| 604 | ||
| 605 | subsection{*Primary Goals of Cardholder Registration *}
 | |
| 606 | ||
| 607 | text{*The cardholder's certificate really was created by the CA, provided the
 | |
| 608 | CA is uncompromised *} | |
| 609 | ||
| 610 | text{*Lemma concerning the actual signed message digest*}
 | |
| 611 | lemma cert_valid_lemma: | |
| 612 |      "[|Crypt (priSK (CA i)) {|Hash {|Nonce N, Pan(pan C)|}, Key cardSK, N1|}
 | |
| 613 | \<in> parts (knows Spy evs); | |
| 614 | CA i \<notin> bad; evs \<in> set_cr|] | |
| 615 | ==> \<exists>KC2 X Y. Says (CA i) C | |
| 616 | (Crypt KC2 | |
| 617 |                        {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
 | |
| 618 | \<in> set evs" | |
| 619 | apply (erule rev_mp) | |
| 620 | apply (erule set_cr.induct) | |
| 621 | apply (simp_all (no_asm_simp)) | |
| 622 | apply auto | |
| 623 | done | |
| 624 | ||
| 625 | text{*Pre-packaged version for cardholder.  We don't try to confirm the values
 | |
| 626 | of KC2, X and Y, since they are not important.*} | |
| 627 | lemma certificate_valid_cardSK: | |
| 628 |     "[|Gets C (Crypt KC2 {|X, certC (pan C) cardSK N onlySig (invKey SKi),
 | |
| 629 | cert (CA i) SKi onlySig (priSK RCA)|}) \<in> set evs; | |
| 630 | CA i \<notin> bad; evs \<in> set_cr|] | |
| 631 | ==> \<exists>KC2 X Y. Says (CA i) C | |
| 632 | (Crypt KC2 | |
| 633 |                        {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
 | |
| 634 | \<in> set evs" | |
| 635 | by (force dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Body] | |
| 636 | certificate_valid_pubSK cert_valid_lemma) | |
| 637 | ||
| 638 | ||
| 639 | lemma Hash_imp_parts [rule_format]: | |
| 640 | "evs \<in> set_cr | |
| 641 |       ==> Hash{|X, Nonce N|} \<in> parts (knows Spy evs) -->
 | |
| 642 | Nonce N \<in> parts (knows Spy evs)" | |
| 643 | apply (erule set_cr.induct, force) | |
| 644 | apply (simp_all (no_asm_simp)) | |
| 645 | apply (blast intro: parts_mono [THEN [2] rev_subsetD]) | |
| 646 | done | |
| 647 | ||
| 648 | lemma Hash_imp_parts2 [rule_format]: | |
| 649 | "evs \<in> set_cr | |
| 650 |       ==> Hash{|X, Nonce M, Y, Nonce N|} \<in> parts (knows Spy evs) -->
 | |
| 651 | Nonce M \<in> parts (knows Spy evs) & Nonce N \<in> parts (knows Spy evs)" | |
| 652 | apply (erule set_cr.induct, force) | |
| 653 | apply (simp_all (no_asm_simp)) | |
| 654 | apply (blast intro: parts_mono [THEN [2] rev_subsetD]) | |
| 655 | done | |
| 656 | ||
| 657 | ||
| 658 | subsection{*Secrecy of Nonces*}
 | |
| 659 | ||
| 660 | subsubsection{*Lemmas about the predicate KeyCryptNonce *}
 | |
| 661 | ||
| 662 | text{*A fresh DK cannot be associated with any other
 | |
| 663 | (with respect to a given trace). *} | |
| 664 | lemma DK_fresh_not_KeyCryptNonce: | |
| 665 | "[| DK \<in> symKeys; Key DK \<notin> used evs; evs \<in> set_cr |] | |
| 666 | ==> ~ KeyCryptNonce DK K evs" | |
| 667 | apply (erule rev_mp) | |
| 668 | apply (erule rev_mp) | |
| 669 | apply (erule set_cr.induct) | |
| 670 | apply (simp_all (no_asm_simp)) | |
| 671 | apply blast | |
| 672 | apply blast | |
| 673 | apply (auto simp add: DK_fresh_not_KeyCryptKey) | |
| 674 | done | |
| 675 | ||
| 676 | text{*A fresh N cannot be associated with any other
 | |
| 677 | (with respect to a given trace). *} | |
| 678 | lemma N_fresh_not_KeyCryptNonce: | |
| 679 | "\<forall>C. DK \<noteq> priEK C ==> Nonce N \<notin> used evs --> ~ KeyCryptNonce DK N evs" | |
| 680 | apply (induct_tac "evs") | |
| 681 | apply (case_tac [2] "a") | |
| 682 | apply (auto simp add: parts_insert2) | |
| 683 | done | |
| 684 | ||
| 685 | lemma not_KeyCryptNonce_cardSK [rule_format (no_asm)]: | |
| 686 | "[|cardSK \<notin> symKeys; \<forall>C. cardSK \<noteq> priEK C; evs \<in> set_cr|] ==> | |
| 687 | Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptNonce cardSK N evs" | |
| 688 | apply (erule set_cr.induct, analz_mono_contra, simp_all) | |
| 14218 | 689 | apply (blast dest: not_KeyCryptKey_cardSK)  --{*6*}
 | 
| 14199 | 690 | done | 
| 691 | ||
| 692 | subsubsection{*Lemmas for message 5 and 6:
 | |
| 693 | either cardSK is compromised (when we don't care) | |
| 694 | or else cardSK hasn't been used to encrypt K. *} | |
| 695 | ||
| 696 | text{*Lemma for message 5: pubSK C is never used to encrypt Nonces.*}
 | |
| 697 | lemma pubSK_not_KeyCryptNonce [simp]: "~ KeyCryptNonce (pubSK C) N evs" | |
| 698 | apply (induct_tac "evs") | |
| 699 | apply (auto simp add: parts_insert2 split add: event.split) | |
| 700 | done | |
| 701 | ||
| 702 | text{*Lemma for message 6: either cardSK is compromised (when we don't care)
 | |
| 703 | or else cardSK hasn't been used to encrypt K.*} | |
| 704 | lemma msg6_KeyCryptNonce_disj: | |
| 705 |      "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
 | |
| 706 | \<in> set evs; | |
| 707 | cardSK \<notin> symKeys; evs \<in> set_cr|] | |
| 708 | ==> Key cardSK \<in> analz (knows Spy evs) | | |
| 709 | ((\<forall>K. ~ KeyCryptKey cardSK K evs) & | |
| 710 | (\<forall>N. ~ KeyCryptNonce cardSK N evs))" | |
| 711 | by (blast dest: not_KeyCryptKey_cardSK not_KeyCryptNonce_cardSK | |
| 712 | intro: cardSK_neq_priEK) | |
| 713 | ||
| 714 | ||
| 715 | text{*As usual: we express the property as a logical equivalence*}
 | |
| 716 | lemma Nonce_analz_image_Key_lemma: | |
| 717 | "P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H) | |
| 718 | ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)" | |
| 719 | by (blast intro: analz_mono [THEN [2] rev_subsetD]) | |
| 720 | ||
| 721 | text{*The @{text "(no_asm)"} attribute is essential, since it retains
 | |
| 722 | the quantifier and allows the simprule's condition to itself be simplified.*} | |
| 723 | lemma Nonce_compromise [rule_format (no_asm)]: | |
| 724 | "evs \<in> set_cr ==> | |
| 725 | (\<forall>N KK. (\<forall>K \<in> KK. ~ KeyCryptNonce K N evs) --> | |
| 726 | (Nonce N \<in> analz (Key`KK Un (knows Spy evs))) = | |
| 727 | (Nonce N \<in> analz (knows Spy evs)))" | |
| 728 | apply (erule set_cr.induct) | |
| 729 | apply (rule_tac [!] allI)+ | |
| 730 | apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+ | |
| 731 | apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
 | |
| 732 | apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
 | |
| 733 | apply (frule_tac [11] msg6_KeyCryptNonce_disj) | |
| 734 | apply (erule_tac [13] disjE) | |
| 735 | apply (simp_all del: image_insert image_Un | |
| 736 | add: symKey_compromise | |
| 737 | analz_image_keys_simps analz_knows_absorb | |
| 738 | analz_Key_image_insert_eq notin_image_iff | |
| 739 | N_fresh_not_KeyCryptNonce | |
| 740 | DK_fresh_not_KeyCryptNonce K_fresh_not_KeyCryptKey | |
| 741 | ball_conj_distrib analz_image_priEK) | |
| 742 |   --{*71 seconds on a 1.8GHz machine*}
 | |
| 14218 | 743 | apply spy_analz  --{*Fake*}
 | 
| 744 | apply blast  --{*3*}
 | |
| 745 | apply blast  --{*5*}
 | |
| 746 | txt{*Message 6*}
 | |
| 14199 | 747 | apply (force del: allE ballE impCE simp add: symKey_compromise) | 
| 14218 | 748 |   --{*cardSK compromised*}
 | 
| 14199 | 749 | txt{*Simplify again--necessary because the previous simplification introduces
 | 
| 750 | some logical connectives*} | |
| 14218 | 751 | apply (force del: allE ballE impCE | 
| 14199 | 752 | simp del: image_insert image_Un imp_disjL | 
| 753 | simp add: analz_image_keys_simps symKey_compromise) | |
| 14218 | 754 | done | 
| 14199 | 755 | |
| 756 | ||
| 757 | subsection{*Secrecy of CardSecret: the Cardholder's secret*}
 | |
| 758 | ||
| 759 | lemma NC2_not_CardSecret: | |
| 760 |      "[|Crypt EKj {|Key K, Pan p, Hash {|Agent D, Nonce N|}|}
 | |
| 761 | \<in> parts (knows Spy evs); | |
| 762 | Key K \<notin> analz (knows Spy evs); | |
| 763 | Nonce N \<notin> analz (knows Spy evs); | |
| 764 | evs \<in> set_cr|] | |
| 765 |       ==> Crypt EKi {|Key K', Pan p', Nonce N|} \<notin> parts (knows Spy evs)"
 | |
| 766 | apply (erule rev_mp) | |
| 767 | apply (erule rev_mp) | |
| 768 | apply (erule rev_mp) | |
| 769 | apply (erule set_cr.induct, analz_mono_contra, simp_all) | |
| 770 | apply (blast dest: Hash_imp_parts)+ | |
| 771 | done | |
| 772 | ||
| 773 | lemma KC2_secure_lemma [rule_format]: | |
| 774 |      "[|U = Crypt KC3 {|Agent C, Nonce N, Key KC2, X|};
 | |
| 775 | U \<in> parts (knows Spy evs); | |
| 776 | evs \<in> set_cr|] | |
| 777 | ==> Nonce N \<notin> analz (knows Spy evs) --> | |
| 778 |       (\<exists>k i W. Says (Cardholder k) (CA i) {|U,W|} \<in> set evs & 
 | |
| 779 | Cardholder k \<notin> bad & CA i \<notin> bad)" | |
| 780 | apply (erule_tac P = "U \<in> ?H" in rev_mp) | |
| 781 | apply (erule set_cr.induct) | |
| 782 | apply (tactic{*valid_certificate_tac 8*})  --{*for message 5*}
 | |
| 783 | apply (simp_all del: image_insert image_Un imp_disjL | |
| 784 | add: analz_image_keys_simps analz_knows_absorb | |
| 785 | analz_knows_absorb2 notin_image_iff) | |
| 786 |   --{*19 seconds on a 1.8GHz machine*}
 | |
| 787 | apply (simp_all (no_asm_simp)) --{*leaves 4 subgoals*}
 | |
| 788 | apply (blast intro!: analz_insertI)+ | |
| 789 | done | |
| 790 | ||
| 791 | lemma KC2_secrecy: | |
| 792 |      "[|Gets B {|Crypt K {|Agent C, Nonce N, Key KC2, X|}, Y|} \<in> set evs;
 | |
| 793 | Nonce N \<notin> analz (knows Spy evs); KC2 \<in> symKeys; | |
| 794 | evs \<in> set_cr|] | |
| 795 | ==> Key KC2 \<notin> analz (knows Spy evs)" | |
| 796 | by (force dest!: refl [THEN KC2_secure_lemma] symKey_secrecy) | |
| 797 | ||
| 798 | ||
| 799 | text{*Inductive version*}
 | |
| 800 | lemma CardSecret_secrecy_lemma [rule_format]: | |
| 801 | "[|CA i \<notin> bad; evs \<in> set_cr|] | |
| 802 | ==> Key K \<notin> analz (knows Spy evs) --> | |
| 803 |           Crypt (pubEK (CA i)) {|Key K, Pan p, Nonce CardSecret|}
 | |
| 804 | \<in> parts (knows Spy evs) --> | |
| 805 | Nonce CardSecret \<notin> analz (knows Spy evs)" | |
| 806 | apply (erule set_cr.induct, analz_mono_contra) | |
| 807 | apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
 | |
| 808 | apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
 | |
| 809 | apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE]) | |
| 810 | apply (simp_all | |
| 811 | del: image_insert image_Un imp_disjL | |
| 812 | add: analz_image_keys_simps analz_knows_absorb | |
| 813 | analz_Key_image_insert_eq notin_image_iff | |
| 814 | EXHcrypt_def Crypt_notin_image_Key | |
| 815 | N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce | |
| 816 | ball_conj_distrib Nonce_compromise symKey_compromise | |
| 817 | analz_image_priEK) | |
| 818 |   --{*12 seconds on a 1.8GHz machine*}
 | |
| 14218 | 819 | apply spy_analz  --{*Fake*}
 | 
| 14199 | 820 | apply (simp_all (no_asm_simp)) | 
| 14218 | 821 | apply blast  --{*1*}
 | 
| 822 | apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])  --{*2*}
 | |
| 823 | apply blast  --{*3*}
 | |
| 824 | apply (blast dest: NC2_not_CardSecret Gets_imp_knows_Spy [THEN analz.Inj] analz_symKeys_Decrypt)  --{*4*}
 | |
| 825 | apply blast  --{*5*}
 | |
| 826 | apply (blast dest: KC2_secrecy)+  --{*Message 6: two cases*}
 | |
| 14199 | 827 | done | 
| 828 | ||
| 829 | ||
| 830 | text{*Packaged version for cardholder*}
 | |
| 831 | lemma CardSecret_secrecy: | |
| 832 | "[|Cardholder k \<notin> bad; CA i \<notin> bad; | |
| 833 | Says (Cardholder k) (CA i) | |
| 834 |            {|X, Crypt EKi {|Key KC3, Pan p, Nonce CardSecret|}|} \<in> set evs;
 | |
| 835 |         Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
 | |
| 836 | cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs; | |
| 837 | KC3 \<in> symKeys; evs \<in> set_cr|] | |
| 838 | ==> Nonce CardSecret \<notin> analz (knows Spy evs)" | |
| 839 | apply (frule Gets_certificate_valid, assumption) | |
| 840 | apply (subgoal_tac "Key KC3 \<notin> analz (knows Spy evs) ") | |
| 841 | apply (blast dest: CardSecret_secrecy_lemma) | |
| 842 | apply (rule symKey_secrecy) | |
| 843 | apply (auto simp add: parts_insert2) | |
| 844 | done | |
| 845 | ||
| 846 | ||
| 847 | subsection{*Secrecy of NonceCCA [the CA's secret] *}
 | |
| 848 | ||
| 849 | lemma NC2_not_NonceCCA: | |
| 850 |      "[|Hash {|Agent C', Nonce N', Agent C, Nonce N|}
 | |
| 851 | \<in> parts (knows Spy evs); | |
| 852 | Nonce N \<notin> analz (knows Spy evs); | |
| 853 | evs \<in> set_cr|] | |
| 854 |       ==> Crypt KC1 {|{|Agent B, Nonce N|}, Hash p|} \<notin> parts (knows Spy evs)"
 | |
| 855 | apply (erule rev_mp) | |
| 856 | apply (erule rev_mp) | |
| 857 | apply (erule set_cr.induct, analz_mono_contra, simp_all) | |
| 858 | apply (blast dest: Hash_imp_parts2)+ | |
| 859 | done | |
| 860 | ||
| 861 | ||
| 862 | text{*Inductive version*}
 | |
| 863 | lemma NonceCCA_secrecy_lemma [rule_format]: | |
| 864 | "[|CA i \<notin> bad; evs \<in> set_cr|] | |
| 865 | ==> Key K \<notin> analz (knows Spy evs) --> | |
| 866 | Crypt K | |
| 867 |             {|sign (priSK (CA i))
 | |
| 868 |                    {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
 | |
| 869 | X, Y|} | |
| 870 | \<in> parts (knows Spy evs) --> | |
| 871 | Nonce NonceCCA \<notin> analz (knows Spy evs)" | |
| 872 | apply (erule set_cr.induct, analz_mono_contra) | |
| 873 | apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
 | |
| 874 | apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
 | |
| 875 | apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE]) | |
| 876 | apply (simp_all | |
| 877 | del: image_insert image_Un imp_disjL | |
| 878 | add: analz_image_keys_simps analz_knows_absorb sign_def | |
| 879 | analz_Key_image_insert_eq notin_image_iff | |
| 880 | EXHcrypt_def Crypt_notin_image_Key | |
| 881 | N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce | |
| 882 | ball_conj_distrib Nonce_compromise symKey_compromise | |
| 883 | analz_image_priEK) | |
| 884 |   --{*15 seconds on a 1.8GHz machine*}
 | |
| 14218 | 885 | apply spy_analz  --{*Fake*}
 | 
| 886 | apply blast  --{*1*}
 | |
| 887 | apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])  --{*2*}
 | |
| 888 | apply blast  --{*3*}
 | |
| 889 | apply (blast dest: NC2_not_NonceCCA)  --{*4*}
 | |
| 890 | apply blast  --{*5*}
 | |
| 891 | apply (blast dest: KC2_secrecy)+  --{*Message 6: two cases*}
 | |
| 14199 | 892 | done | 
| 893 | ||
| 894 | ||
| 895 | text{*Packaged version for cardholder*}
 | |
| 896 | lemma NonceCCA_secrecy: | |
| 897 | "[|Cardholder k \<notin> bad; CA i \<notin> bad; | |
| 898 | Gets (Cardholder k) | |
| 899 | (Crypt KC2 | |
| 900 |             {|sign (priSK (CA i)) {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
 | |
| 901 | X, Y|}) \<in> set evs; | |
| 902 | Says (Cardholder k) (CA i) | |
| 903 |            {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, X'|}, Y'|} \<in> set evs;
 | |
| 904 |         Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
 | |
| 905 | cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs; | |
| 906 | KC2 \<in> symKeys; evs \<in> set_cr|] | |
| 907 | ==> Nonce NonceCCA \<notin> analz (knows Spy evs)" | |
| 908 | apply (frule Gets_certificate_valid, assumption) | |
| 909 | apply (subgoal_tac "Key KC2 \<notin> analz (knows Spy evs) ") | |
| 910 | apply (blast dest: NonceCCA_secrecy_lemma) | |
| 911 | apply (rule symKey_secrecy) | |
| 912 | apply (auto simp add: parts_insert2) | |
| 913 | done | |
| 914 | ||
| 915 | text{*We don't bother to prove guarantees for the CA.  He doesn't care about
 | |
| 916 | the PANSecret: it isn't his credit card!*} | |
| 917 | ||
| 918 | ||
| 919 | subsection{*Rewriting Rule for PANs*}
 | |
| 920 | ||
| 921 | text{*Lemma for message 6: either cardSK isn't a CA's private encryption key,
 | |
| 922 | or if it is then (because it appears in traffic) that CA is bad, | |
| 923 | and so the Spy knows that key already. Either way, we can simplify | |
| 924 |   the expression @{term "analz (insert (Key cardSK) X)"}.*}
 | |
| 925 | lemma msg6_cardSK_disj: | |
| 926 |      "[|Gets A {|Crypt K {|c, n, k', Key cardSK, X|}, Y|}
 | |
| 927 | \<in> set evs; evs \<in> set_cr |] | |
| 928 | ==> cardSK \<notin> range(invKey o pubEK o CA) | Key cardSK \<in> knows Spy evs" | |
| 929 | by auto | |
| 930 | ||
| 931 | lemma analz_image_pan_lemma: | |
| 932 | "(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H) ==> | |
| 933 | (Pan P \<in> analz (Key`nE Un H)) = (Pan P \<in> analz H)" | |
| 934 | by (blast intro: analz_mono [THEN [2] rev_subsetD]) | |
| 935 | ||
| 936 | lemma analz_image_pan [rule_format]: | |
| 937 | "evs \<in> set_cr ==> | |
| 938 | \<forall>KK. KK <= - invKey ` pubEK ` range CA --> | |
| 939 | (Pan P \<in> analz (Key`KK Un (knows Spy evs))) = | |
| 940 | (Pan P \<in> analz (knows Spy evs))" | |
| 941 | apply (erule set_cr.induct) | |
| 942 | apply (rule_tac [!] allI impI)+ | |
| 943 | apply (rule_tac [!] analz_image_pan_lemma) | |
| 944 | apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
 | |
| 945 | apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
 | |
| 946 | apply (erule_tac [9] msg6_cardSK_disj [THEN disjE]) | |
| 947 | apply (simp_all | |
| 948 | del: image_insert image_Un | |
| 949 | add: analz_image_keys_simps disjoint_image_iff | |
| 950 | notin_image_iff analz_image_priEK) | |
| 951 |   --{*33 seconds on a 1.8GHz machine*}
 | |
| 952 | apply spy_analz | |
| 14218 | 953 | apply (simp add: insert_absorb)  --{*6*}
 | 
| 14199 | 954 | done | 
| 955 | ||
| 956 | lemma analz_insert_pan: | |
| 957 | "[| evs \<in> set_cr; K \<notin> invKey ` pubEK ` range CA |] ==> | |
| 958 | (Pan P \<in> analz (insert (Key K) (knows Spy evs))) = | |
| 959 | (Pan P \<in> analz (knows Spy evs))" | |
| 960 | by (simp del: image_insert image_Un | |
| 961 | add: analz_image_keys_simps analz_image_pan) | |
| 962 | ||
| 963 | ||
| 964 | text{*Confidentiality of the PAN\@.  Maybe we could combine the statements of
 | |
| 965 |   this theorem with @{term analz_image_pan}, requiring a single induction but
 | |
| 966 | a much more difficult proof.*} | |
| 967 | lemma pan_confidentiality: | |
| 968 | "[| Pan (pan C) \<in> analz(knows Spy evs); C \<noteq>Spy; evs :set_cr|] | |
| 969 | ==> \<exists>i X K HN. | |
| 970 |         Says C (CA i) {|X, Crypt (pubEK (CA i)) {|Key K, Pan (pan C), HN|} |}
 | |
| 971 | \<in> set evs | |
| 972 | & (CA i) \<in> bad" | |
| 973 | apply (erule rev_mp) | |
| 974 | apply (erule set_cr.induct) | |
| 975 | apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
 | |
| 976 | apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
 | |
| 977 | apply (erule_tac [9] msg6_cardSK_disj [THEN disjE]) | |
| 978 | apply (simp_all | |
| 979 | del: image_insert image_Un | |
| 980 | add: analz_image_keys_simps analz_insert_pan analz_image_pan | |
| 981 | notin_image_iff analz_image_priEK) | |
| 982 |   --{*18 seconds on a 1.8GHz machine*}
 | |
| 14218 | 983 | apply spy_analz  --{*fake*}
 | 
| 984 | apply blast  --{*3*}
 | |
| 985 | apply blast  --{*5*}
 | |
| 986 | apply (simp (no_asm_simp) add: insert_absorb)  --{*6*}
 | |
| 14199 | 987 | done | 
| 988 | ||
| 989 | ||
| 990 | subsection{*Unicity*}
 | |
| 991 | ||
| 992 | lemma CR6_Says_imp_Notes: | |
| 993 | "[|Says (CA i) C (Crypt KC2 | |
| 994 |           {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
 | |
| 995 | certC (pan C) cardSK X onlySig (priSK (CA i)), | |
| 996 | cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}) \<in> set evs; | |
| 997 | evs \<in> set_cr |] | |
| 998 | ==> Notes (CA i) (Key cardSK) \<in> set evs" | |
| 999 | apply (erule rev_mp) | |
| 1000 | apply (erule set_cr.induct) | |
| 1001 | apply (simp_all (no_asm_simp)) | |
| 1002 | done | |
| 1003 | ||
| 1004 | text{*Unicity of cardSK: it uniquely identifies the other components.  
 | |
| 1005 | This holds because a CA accepts a cardSK at most once.*} | |
| 1006 | lemma cardholder_key_unicity: | |
| 1007 | "[|Says (CA i) C (Crypt KC2 | |
| 1008 |           {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
 | |
| 1009 | certC (pan C) cardSK X onlySig (priSK (CA i)), | |
| 1010 | cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}) | |
| 1011 | \<in> set evs; | |
| 1012 | Says (CA i) C' (Crypt KC2' | |
| 1013 |           {|sign (priSK (CA i)) {|Agent C', Nonce NC3', Agent (CA i), Nonce Y'|},
 | |
| 1014 | certC (pan C') cardSK X' onlySig (priSK (CA i)), | |
| 1015 | cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}) | |
| 1016 | \<in> set evs; | |
| 1017 | evs \<in> set_cr |] ==> C=C' & NC3=NC3' & X=X' & KC2=KC2' & Y=Y'" | |
| 1018 | apply (erule rev_mp) | |
| 1019 | apply (erule rev_mp) | |
| 1020 | apply (erule set_cr.induct) | |
| 1021 | apply (simp_all (no_asm_simp)) | |
| 1022 | apply (blast dest!: CR6_Says_imp_Notes) | |
| 1023 | done | |
| 1024 | ||
| 1025 | ||
| 14218 | 1026 | (*<*) | 
| 14199 | 1027 | text{*UNUSED unicity result*}
 | 
| 1028 | lemma unique_KC1: | |
| 1029 |      "[|Says C B {|Crypt KC1 X, Crypt EK {|Key KC1, Y|}|}
 | |
| 1030 | \<in> set evs; | |
| 1031 |         Says C B' {|Crypt KC1 X', Crypt EK' {|Key KC1, Y'|}|}
 | |
| 1032 | \<in> set evs; | |
| 1033 | C \<notin> bad; evs \<in> set_cr|] ==> B'=B & Y'=Y" | |
| 1034 | apply (erule rev_mp) | |
| 1035 | apply (erule rev_mp) | |
| 1036 | apply (erule set_cr.induct, auto) | |
| 1037 | done | |
| 1038 | ||
| 1039 | text{*UNUSED unicity result*}
 | |
| 1040 | lemma unique_KC2: | |
| 1041 |      "[|Says C B {|Crypt K {|Agent C, nn, Key KC2, X|}, Y|} \<in> set evs;
 | |
| 1042 |         Says C B' {|Crypt K' {|Agent C, nn', Key KC2, X'|}, Y'|} \<in> set evs;
 | |
| 1043 | C \<notin> bad; evs \<in> set_cr|] ==> B'=B & X'=X" | |
| 1044 | apply (erule rev_mp) | |
| 1045 | apply (erule rev_mp) | |
| 1046 | apply (erule set_cr.induct, auto) | |
| 1047 | done | |
| 14218 | 1048 | (*>*) | 
| 1049 | ||
| 14199 | 1050 | |
| 1051 | text{*Cannot show cardSK to be secret because it isn't assumed to be fresh
 | |
| 1052 | it could be a previously compromised cardSK [e.g. involving a bad CA]*} | |
| 1053 | ||
| 1054 | ||
| 1055 | end |