| author | huffman | 
| Thu, 27 Mar 2008 00:27:16 +0100 | |
| changeset 26420 | 57a626f64875 | 
| parent 24123 | a0fc58900606 | 
| child 30510 | 4120fc59dd85 | 
| permissions | -rw-r--r-- | 
| 14199 | 1 | (* Title: HOL/Auth/SET/PublicSET | 
| 2 | ID: $Id$ | |
| 3 | Authors: Giampaolo Bella, Fabio Massacci, Lawrence C Paulson | |
| 4 | *) | |
| 5 | ||
| 6 | header{*The Public-Key Theory, Modified for SET*}
 | |
| 7 | ||
| 16417 | 8 | theory PublicSET imports EventSET begin | 
| 14199 | 9 | |
| 10 | subsection{*Symmetric and Asymmetric Keys*}
 | |
| 11 | ||
| 12 | text{*definitions influenced by the wish to assign asymmetric keys 
 | |
| 13 | - since the beginning - only to RCA and CAs, namely we need a partial | |
| 14 | function on type Agent*} | |
| 15 | ||
| 16 | ||
| 17 | text{*The SET specs mention two signature keys for CAs - we only have one*}
 | |
| 18 | ||
| 19 | consts | |
| 20 | publicKey :: "[bool, agent] => key" | |
| 21 |     --{*the boolean is TRUE if a signing key*}
 | |
| 22 | ||
| 23 | syntax | |
| 24 | pubEK :: "agent => key" | |
| 25 | pubSK :: "agent => key" | |
| 26 | priEK :: "agent => key" | |
| 27 | priSK :: "agent => key" | |
| 28 | ||
| 29 | translations | |
| 30 | "pubEK" == "publicKey False" | |
| 31 | "pubSK" == "publicKey True" | |
| 32 | ||
| 33 | (*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*) | |
| 34 | "priEK A" == "invKey (pubEK A)" | |
| 35 | "priSK A" == "invKey (pubSK A)" | |
| 36 | ||
| 37 | text{*By freeness of agents, no two agents have the same key. Since
 | |
| 38 |  @{term "True\<noteq>False"}, no agent has the same signing and encryption keys.*}
 | |
| 39 | ||
| 40 | specification (publicKey) | |
| 41 | injective_publicKey: | |
| 42 | "publicKey b A = publicKey c A' ==> b=c & A=A'" | |
| 14218 | 43 | (*<*) | 
| 14199 | 44 | apply (rule exI [of _ "%b A. 2 * nat_of_agent A + (if b then 1 else 0)"]) | 
| 45 | apply (auto simp add: inj_on_def inj_nat_of_agent [THEN inj_eq] split: agent.split) | |
| 46 | apply (drule_tac f="%x. x mod 2" in arg_cong, simp add: mod_Suc)+ | |
| 47 | (*or this, but presburger won't abstract out the function applications | |
| 48 | apply presburger+ | |
| 49 | *) | |
| 50 | done | |
| 14218 | 51 | (*>*) | 
| 14199 | 52 | |
| 53 | axioms | |
| 54 | (*No private key equals any public key (essential to ensure that private | |
| 55 | keys are private!) *) | |
| 56 | privateKey_neq_publicKey [iff]: | |
| 57 | "invKey (publicKey b A) \<noteq> publicKey b' A'" | |
| 58 | ||
| 59 | declare privateKey_neq_publicKey [THEN not_sym, iff] | |
| 60 | ||
| 61 | ||
| 62 | subsection{*Initial Knowledge*}
 | |
| 63 | ||
| 64 | text{*This information is not necessary.  Each protocol distributes any needed
 | |
| 65 | certificates, and anyway our proofs require a formalization of the Spy's | |
| 66 | knowledge only. However, the initial knowledge is as follows: | |
| 67 | All agents know RCA's public keys; | |
| 68 | RCA and CAs know their own respective keys; | |
| 69 | RCA (has already certified and therefore) knows all CAs public keys; | |
| 70 | Spy knows all keys of all bad agents.*} | |
| 71 | primrec | |
| 14218 | 72 | (*<*) | 
| 14199 | 73 | initState_CA: | 
| 74 | "initState (CA i) = | |
| 75 |        (if i=0 then Key ` ({priEK RCA, priSK RCA} Un
 | |
| 76 | pubEK ` (range CA) Un pubSK ` (range CA)) | |
| 77 | 	else {Key (priEK (CA i)), Key (priSK (CA i)),
 | |
| 78 | Key (pubEK (CA i)), Key (pubSK (CA i)), | |
| 79 | Key (pubEK RCA), Key (pubSK RCA)})" | |
| 80 | ||
| 81 | initState_Cardholder: | |
| 82 | "initState (Cardholder i) = | |
| 83 |        {Key (priEK (Cardholder i)), Key (priSK (Cardholder i)),
 | |
| 84 | Key (pubEK (Cardholder i)), Key (pubSK (Cardholder i)), | |
| 85 | Key (pubEK RCA), Key (pubSK RCA)}" | |
| 86 | ||
| 87 | initState_Merchant: | |
| 88 | "initState (Merchant i) = | |
| 89 |        {Key (priEK (Merchant i)), Key (priSK (Merchant i)),
 | |
| 90 | Key (pubEK (Merchant i)), Key (pubSK (Merchant i)), | |
| 91 | Key (pubEK RCA), Key (pubSK RCA)}" | |
| 92 | ||
| 93 | initState_PG: | |
| 94 | "initState (PG i) = | |
| 95 |        {Key (priEK (PG i)), Key (priSK (PG i)),
 | |
| 96 | Key (pubEK (PG i)), Key (pubSK (PG i)), | |
| 97 | Key (pubEK RCA), Key (pubSK RCA)}" | |
| 14218 | 98 | (*>*) | 
| 14199 | 99 | initState_Spy: | 
| 100 | "initState Spy = Key ` (invKey ` pubEK ` bad Un | |
| 101 | invKey ` pubSK ` bad Un | |
| 102 | range pubEK Un range pubSK)" | |
| 103 | ||
| 104 | ||
| 105 | text{*Injective mapping from agents to PANs: an agent can have only one card*}
 | |
| 106 | ||
| 107 | consts pan :: "agent => nat" | |
| 108 | ||
| 109 | specification (pan) | |
| 110 | inj_pan: "inj pan" | |
| 111 |   --{*No two agents have the same PAN*}
 | |
| 14218 | 112 | (*<*) | 
| 14199 | 113 | apply (rule exI [of _ "nat_of_agent"]) | 
| 114 | apply (simp add: inj_on_def inj_nat_of_agent [THEN inj_eq]) | |
| 115 | done | |
| 14218 | 116 | (*>*) | 
| 14199 | 117 | |
| 118 | declare inj_pan [THEN inj_eq, iff] | |
| 119 | ||
| 120 | consts | |
| 121 |   XOR :: "nat*nat => nat"  --{*no properties are assumed of exclusive-or*}
 | |
| 122 | ||
| 123 | ||
| 124 | subsection{*Signature Primitives*}
 | |
| 125 | ||
| 126 | constdefs | |
| 127 | ||
| 128 | (* Signature = Message + signed Digest *) | |
| 129 | sign :: "[key, msg]=>msg" | |
| 130 |     "sign K X == {|X, Crypt K (Hash X) |}"
 | |
| 131 | ||
| 132 | (* Signature Only = signed Digest Only *) | |
| 133 | signOnly :: "[key, msg]=>msg" | |
| 134 | "signOnly K X == Crypt K (Hash X)" | |
| 135 | ||
| 136 | (* Signature for Certificates = Message + signed Message *) | |
| 137 | signCert :: "[key, msg]=>msg" | |
| 138 |     "signCert K X == {|X, Crypt K X |}"
 | |
| 139 | ||
| 140 | (* Certification Authority's Certificate. | |
| 141 | Contains agent name, a key, a number specifying the key's target use, | |
| 142 | a key to sign the entire certificate. | |
| 143 | ||
| 144 | Should prove if signK=priSK RCA and C=CA i, | |
| 145 | then Ka=pubEK i or pubSK i depending on T ?? | |
| 146 | *) | |
| 147 | cert :: "[agent, key, msg, key] => msg" | |
| 148 |     "cert A Ka T signK == signCert signK {|Agent A, Key Ka, T|}"
 | |
| 149 | ||
| 150 | ||
| 151 | (* Cardholder's Certificate. | |
| 152 | Contains a PAN, the certified key Ka, the PANSecret PS, | |
| 153 | a number specifying the target use for Ka, the signing key signK. | |
| 154 | *) | |
| 155 | certC :: "[nat, key, nat, msg, key] => msg" | |
| 156 | "certC PAN Ka PS T signK == | |
| 157 |      signCert signK {|Hash {|Nonce PS, Pan PAN|}, Key Ka, T|}"
 | |
| 158 | ||
| 159 | (*cert and certA have no repeated elements, so they could be translations, | |
| 160 | but that's tricky and makes proofs slower*) | |
| 161 | ||
| 162 | syntax | |
| 163 | "onlyEnc" :: msg | |
| 164 | "onlySig" :: msg | |
| 165 | "authCode" :: msg | |
| 166 | ||
| 167 | translations | |
| 168 | "onlyEnc" == "Number 0" | |
| 169 | "onlySig" == "Number (Suc 0)" | |
| 170 | "authCode" == "Number (Suc (Suc 0))" | |
| 171 | ||
| 172 | subsection{*Encryption Primitives*}
 | |
| 173 | ||
| 174 | constdefs | |
| 175 | ||
| 176 | EXcrypt :: "[key,key,msg,msg] => msg" | |
| 177 |   --{*Extra Encryption*}
 | |
| 178 | (*K: the symmetric key EK: the public encryption key*) | |
| 179 | "EXcrypt K EK M m == | |
| 180 |        {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m|}|}"
 | |
| 181 | ||
| 182 | EXHcrypt :: "[key,key,msg,msg] => msg" | |
| 183 |   --{*Extra Encryption with Hashing*}
 | |
| 184 | (*K: the symmetric key EK: the public encryption key*) | |
| 185 | "EXHcrypt K EK M m == | |
| 186 |        {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m, Hash M|}|}"
 | |
| 187 | ||
| 188 | Enc :: "[key,key,key,msg] => msg" | |
| 189 |   --{*Simple Encapsulation with SIGNATURE*}
 | |
| 190 | (*SK: the sender's signing key | |
| 191 | K: the symmetric key | |
| 192 | EK: the public encryption key*) | |
| 193 | "Enc SK K EK M == | |
| 194 |        {|Crypt K (sign SK M), Crypt EK (Key K)|}"
 | |
| 195 | ||
| 196 | EncB :: "[key,key,key,msg,msg] => msg" | |
| 197 |   --{*Encapsulation with Baggage.  Keys as above, and baggage b.*}
 | |
| 198 | "EncB SK K EK M b == | |
| 199 |        {|Enc SK K EK {|M, Hash b|}, b|}"
 | |
| 200 | ||
| 201 | ||
| 202 | subsection{*Basic Properties of pubEK, pubSK, priEK and priSK *}
 | |
| 203 | ||
| 204 | lemma publicKey_eq_iff [iff]: | |
| 205 | "(publicKey b A = publicKey b' A') = (b=b' & A=A')" | |
| 206 | by (blast dest: injective_publicKey) | |
| 207 | ||
| 208 | lemma privateKey_eq_iff [iff]: | |
| 209 | "(invKey (publicKey b A) = invKey (publicKey b' A')) = (b=b' & A=A')" | |
| 210 | by auto | |
| 211 | ||
| 212 | lemma not_symKeys_publicKey [iff]: "publicKey b A \<notin> symKeys" | |
| 213 | by (simp add: symKeys_def) | |
| 214 | ||
| 215 | lemma not_symKeys_privateKey [iff]: "invKey (publicKey b A) \<notin> symKeys" | |
| 216 | by (simp add: symKeys_def) | |
| 217 | ||
| 218 | lemma symKeys_invKey_eq [simp]: "K \<in> symKeys ==> invKey K = K" | |
| 219 | by (simp add: symKeys_def) | |
| 220 | ||
| 221 | lemma symKeys_invKey_iff [simp]: "(invKey K \<in> symKeys) = (K \<in> symKeys)" | |
| 222 | by (unfold symKeys_def, auto) | |
| 223 | ||
| 224 | text{*Can be slow (or even loop) as a simprule*}
 | |
| 225 | lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'" | |
| 226 | by blast | |
| 227 | ||
| 228 | text{*These alternatives to @{text symKeys_neq_imp_neq} don't seem any better
 | |
| 229 | in practice.*} | |
| 230 | lemma publicKey_neq_symKey: "K \<in> symKeys ==> publicKey b A \<noteq> K" | |
| 231 | by blast | |
| 232 | ||
| 233 | lemma symKey_neq_publicKey: "K \<in> symKeys ==> K \<noteq> publicKey b A" | |
| 234 | by blast | |
| 235 | ||
| 236 | lemma privateKey_neq_symKey: "K \<in> symKeys ==> invKey (publicKey b A) \<noteq> K" | |
| 237 | by blast | |
| 238 | ||
| 239 | lemma symKey_neq_privateKey: "K \<in> symKeys ==> K \<noteq> invKey (publicKey b A)" | |
| 240 | by blast | |
| 241 | ||
| 242 | lemma analz_symKeys_Decrypt: | |
| 243 | "[| Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H |] | |
| 244 | ==> X \<in> analz H" | |
| 245 | by auto | |
| 246 | ||
| 247 | ||
| 248 | subsection{*"Image" Equations That Hold for Injective Functions *}
 | |
| 249 | ||
| 250 | lemma invKey_image_eq [iff]: "(invKey x \<in> invKey`A) = (x\<in>A)" | |
| 251 | by auto | |
| 252 | ||
| 253 | text{*holds because invKey is injective*}
 | |
| 254 | lemma publicKey_image_eq [iff]: | |
| 255 | "(publicKey b A \<in> publicKey c ` AS) = (b=c & A\<in>AS)" | |
| 256 | by auto | |
| 257 | ||
| 258 | lemma privateKey_image_eq [iff]: | |
| 259 | "(invKey (publicKey b A) \<in> invKey ` publicKey c ` AS) = (b=c & A\<in>AS)" | |
| 260 | by auto | |
| 261 | ||
| 262 | lemma privateKey_notin_image_publicKey [iff]: | |
| 263 | "invKey (publicKey b A) \<notin> publicKey c ` AS" | |
| 264 | by auto | |
| 265 | ||
| 266 | lemma publicKey_notin_image_privateKey [iff]: | |
| 267 | "publicKey b A \<notin> invKey ` publicKey c ` AS" | |
| 268 | by auto | |
| 269 | ||
| 270 | lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
 | |
| 271 | apply (simp add: keysFor_def) | |
| 272 | apply (induct_tac "C") | |
| 273 | apply (auto intro: range_eqI) | |
| 274 | done | |
| 275 | ||
| 276 | text{*for proving @{text new_keys_not_used}*}
 | |
| 277 | lemma keysFor_parts_insert: | |
| 278 | "[| K \<in> keysFor (parts (insert X H)); X \<in> synth (analz H) |] | |
| 279 | ==> K \<in> keysFor (parts H) | Key (invKey K) \<in> parts H" | |
| 14218 | 280 | by (force dest!: | 
| 14199 | 281 | parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD] | 
| 282 | analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD] | |
| 283 | intro: analz_into_parts) | |
| 284 | ||
| 285 | lemma Crypt_imp_keysFor [intro]: | |
| 286 | "[|K \<in> symKeys; Crypt K X \<in> H|] ==> K \<in> keysFor H" | |
| 287 | by (drule Crypt_imp_invKey_keysFor, simp) | |
| 288 | ||
| 289 | text{*Agents see their own private keys!*}
 | |
| 290 | lemma privateKey_in_initStateCA [iff]: | |
| 291 | "Key (invKey (publicKey b A)) \<in> initState A" | |
| 292 | by (case_tac "A", auto) | |
| 293 | ||
| 294 | text{*Agents see their own public keys!*}
 | |
| 295 | lemma publicKey_in_initStateCA [iff]: "Key (publicKey b A) \<in> initState A" | |
| 296 | by (case_tac "A", auto) | |
| 297 | ||
| 298 | text{*RCA sees CAs' public keys! *}
 | |
| 299 | lemma pubK_CA_in_initState_RCA [iff]: | |
| 300 | "Key (publicKey b (CA i)) \<in> initState RCA" | |
| 301 | by auto | |
| 302 | ||
| 303 | ||
| 304 | text{*Spy knows all public keys*}
 | |
| 305 | lemma knows_Spy_pubEK_i [iff]: "Key (publicKey b A) \<in> knows Spy evs" | |
| 306 | apply (induct_tac "evs") | |
| 307 | apply (simp_all add: imageI knows_Cons split add: event.split) | |
| 308 | done | |
| 309 | ||
| 310 | declare knows_Spy_pubEK_i [THEN analz.Inj, iff] | |
| 311 | (*needed????*) | |
| 312 | ||
| 313 | text{*Spy sees private keys of bad agents! [and obviously public keys too]*}
 | |
| 314 | lemma knows_Spy_bad_privateKey [intro!]: | |
| 315 | "A \<in> bad ==> Key (invKey (publicKey b A)) \<in> knows Spy evs" | |
| 14206 
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
 paulson parents: 
14199diff
changeset | 316 | by (rule initState_subset_knows [THEN subsetD], simp) | 
| 14199 | 317 | |
| 318 | ||
| 319 | subsection{*Fresh Nonces for Possibility Theorems*}
 | |
| 320 | ||
| 321 | lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)" | |
| 322 | by (induct_tac "B", auto) | |
| 323 | ||
| 324 | lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []" | |
| 325 | by (simp add: used_Nil) | |
| 326 | ||
| 327 | text{*In any trace, there is an upper bound N on the greatest nonce in use.*}
 | |
| 328 | lemma Nonce_supply_lemma: "\<exists>N. \<forall>n. N<=n --> Nonce n \<notin> used evs" | |
| 329 | apply (induct_tac "evs") | |
| 330 | apply (rule_tac x = 0 in exI) | |
| 331 | apply (simp_all add: used_Cons split add: event.split, safe) | |
| 332 | apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+ | |
| 333 | done | |
| 334 | ||
| 335 | lemma Nonce_supply1: "\<exists>N. Nonce N \<notin> used evs" | |
| 336 | by (rule Nonce_supply_lemma [THEN exE], blast) | |
| 337 | ||
| 338 | lemma Nonce_supply: "Nonce (@ N. Nonce N \<notin> used evs) \<notin> used evs" | |
| 339 | apply (rule Nonce_supply_lemma [THEN exE]) | |
| 340 | apply (rule someI, fast) | |
| 341 | done | |
| 342 | ||
| 343 | ||
| 344 | subsection{*Specialized Methods for Possibility Theorems*}
 | |
| 345 | ||
| 346 | ML | |
| 347 | {*
 | |
| 24123 | 348 | structure PublicSET = | 
| 349 | struct | |
| 14199 | 350 | |
| 351 | (*Tactic for possibility theorems (Isar interface)*) | |
| 352 | fun gen_possibility_tac ss state = state |> | |
| 353 | REPEAT (*omit used_Says so that Nonces start from different traces!*) | |
| 24123 | 354 |     (ALLGOALS (simp_tac (ss delsimps [@{thm used_Says}, @{thm used_Notes}]))
 | 
| 14199 | 355 | THEN | 
| 356 | REPEAT_FIRST (eq_assume_tac ORELSE' | |
| 24123 | 357 |                    resolve_tac [refl, conjI, @{thm Nonce_supply}]))
 | 
| 14199 | 358 | |
| 359 | (*Tactic for possibility theorems (ML script version)*) | |
| 20048 | 360 | fun possibility_tac state = gen_possibility_tac (simpset_of (Thm.theory_of_thm state)) state | 
| 14199 | 361 | |
| 362 | (*For harder protocols (such as SET_CR!), where we have to set up some | |
| 363 | nonces and keys initially*) | |
| 364 | fun basic_possibility_tac st = st |> | |
| 365 | REPEAT | |
| 20048 | 366 | (ALLGOALS (asm_simp_tac (simpset_of (Thm.theory_of_thm st) setSolver safe_solver)) | 
| 14199 | 367 | THEN | 
| 368 | REPEAT_FIRST (resolve_tac [refl, conjI])) | |
| 24123 | 369 | |
| 370 | end | |
| 14199 | 371 | *} | 
| 372 | ||
| 373 | method_setup possibility = {*
 | |
| 374 | Method.ctxt_args (fn ctxt => | |
| 24123 | 375 | Method.SIMPLE_METHOD (PublicSET.gen_possibility_tac (local_simpset_of ctxt))) *} | 
| 14199 | 376 | "for proving possibility theorems" | 
| 377 | ||
| 378 | ||
| 379 | ||
| 380 | subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
 | |
| 381 | ||
| 382 | lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H"
 | |
| 383 | by blast | |
| 384 | ||
| 385 | lemma insert_Key_image: | |
| 386 | "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C" | |
| 387 | by blast | |
| 388 | ||
| 389 | text{*Needed for @{text DK_fresh_not_KeyCryptKey}*}
 | |
| 390 | lemma publicKey_in_used [iff]: "Key (publicKey b A) \<in> used evs" | |
| 391 | by auto | |
| 392 | ||
| 393 | lemma privateKey_in_used [iff]: "Key (invKey (publicKey b A)) \<in> used evs" | |
| 394 | by (blast intro!: initState_into_used) | |
| 395 | ||
| 396 | text{*Reverse the normal simplification of "image" to build up (not break down)
 | |
| 397 |   the set of keys.  Based on @{text analz_image_freshK_ss}, but simpler.*}
 | |
| 398 | lemmas analz_image_keys_simps = | |
| 399 |        simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
 | |
| 400 | image_insert [THEN sym] image_Un [THEN sym] | |
| 401 | rangeI symKeys_neq_imp_neq | |
| 402 | insert_Key_singleton insert_Key_image Un_assoc [THEN sym] | |
| 403 | ||
| 404 | ||
| 405 | (*General lemmas proved by Larry*) | |
| 406 | ||
| 407 | subsection{*Controlled Unfolding of Abbreviations*}
 | |
| 408 | ||
| 409 | text{*A set is expanded only if a relation is applied to it*}
 | |
| 410 | lemma def_abbrev_simp_relation: | |
| 411 | "A == B ==> (A \<in> X) = (B \<in> X) & | |
| 412 | (u = A) = (u = B) & | |
| 413 | (A = u) = (B = u)" | |
| 414 | by auto | |
| 415 | ||
| 416 | text{*A set is expanded only if one of the given functions is applied to it*}
 | |
| 417 | lemma def_abbrev_simp_function: | |
| 418 | "A == B | |
| 419 | ==> parts (insert A X) = parts (insert B X) & | |
| 420 | analz (insert A X) = analz (insert B X) & | |
| 421 | keysFor (insert A X) = keysFor (insert B X)" | |
| 422 | by auto | |
| 423 | ||
| 424 | subsubsection{*Special Simplification Rules for @{term signCert}*}
 | |
| 425 | ||
| 426 | text{*Avoids duplicating X and its components!*}
 | |
| 427 | lemma parts_insert_signCert: | |
| 428 | "parts (insert (signCert K X) H) = | |
| 429 |       insert {|X, Crypt K X|} (parts (insert (Crypt K X) H))"
 | |
| 430 | by (simp add: signCert_def insert_commute [of X]) | |
| 431 | ||
| 432 | text{*Avoids a case split! [X is always available]*}
 | |
| 433 | lemma analz_insert_signCert: | |
| 434 | "analz (insert (signCert K X) H) = | |
| 435 |       insert {|X, Crypt K X|} (insert (Crypt K X) (analz (insert X H)))"
 | |
| 436 | by (simp add: signCert_def insert_commute [of X]) | |
| 437 | ||
| 438 | lemma keysFor_insert_signCert: "keysFor (insert (signCert K X) H) = keysFor H" | |
| 439 | by (simp add: signCert_def) | |
| 440 | ||
| 441 | text{*Controlled rewrite rules for @{term signCert}, just the definitions
 | |
| 442 | of the others. Encryption primitives are just expanded, despite their huge | |
| 443 | redundancy!*} | |
| 444 | lemmas abbrev_simps [simp] = | |
| 445 | parts_insert_signCert analz_insert_signCert keysFor_insert_signCert | |
| 446 | sign_def [THEN def_abbrev_simp_relation] | |
| 447 | sign_def [THEN def_abbrev_simp_function] | |
| 448 | signCert_def [THEN def_abbrev_simp_relation] | |
| 449 | signCert_def [THEN def_abbrev_simp_function] | |
| 450 | certC_def [THEN def_abbrev_simp_relation] | |
| 451 | certC_def [THEN def_abbrev_simp_function] | |
| 452 | cert_def [THEN def_abbrev_simp_relation] | |
| 453 | cert_def [THEN def_abbrev_simp_function] | |
| 454 | EXcrypt_def [THEN def_abbrev_simp_relation] | |
| 455 | EXcrypt_def [THEN def_abbrev_simp_function] | |
| 456 | EXHcrypt_def [THEN def_abbrev_simp_relation] | |
| 457 | EXHcrypt_def [THEN def_abbrev_simp_function] | |
| 458 | Enc_def [THEN def_abbrev_simp_relation] | |
| 459 | Enc_def [THEN def_abbrev_simp_function] | |
| 460 | EncB_def [THEN def_abbrev_simp_relation] | |
| 461 | EncB_def [THEN def_abbrev_simp_function] | |
| 462 | ||
| 463 | ||
| 464 | subsubsection{*Elimination Rules for Controlled Rewriting *}
 | |
| 465 | ||
| 466 | lemma Enc_partsE: | |
| 467 | "!!R. [|Enc SK K EK M \<in> parts H; | |
| 468 | [|Crypt K (sign SK M) \<in> parts H; | |
| 469 | Crypt EK (Key K) \<in> parts H|] ==> R|] | |
| 470 | ==> R" | |
| 471 | ||
| 472 | by (unfold Enc_def, blast) | |
| 473 | ||
| 474 | lemma EncB_partsE: | |
| 475 | "!!R. [|EncB SK K EK M b \<in> parts H; | |
| 476 |              [|Crypt K (sign SK {|M, Hash b|}) \<in> parts H;  
 | |
| 477 | Crypt EK (Key K) \<in> parts H; | |
| 478 | b \<in> parts H|] ==> R|] | |
| 479 | ==> R" | |
| 480 | by (unfold EncB_def Enc_def, blast) | |
| 481 | ||
| 482 | lemma EXcrypt_partsE: | |
| 483 | "!!R. [|EXcrypt K EK M m \<in> parts H; | |
| 484 |              [|Crypt K {|M, Hash m|} \<in> parts H;  
 | |
| 485 |                Crypt EK {|Key K, m|} \<in> parts H|] ==> R|]  
 | |
| 486 | ==> R" | |
| 487 | by (unfold EXcrypt_def, blast) | |
| 488 | ||
| 489 | ||
| 490 | subsection{*Lemmas to Simplify Expressions Involving @{term analz} *}
 | |
| 491 | ||
| 492 | lemma analz_knows_absorb: | |
| 493 | "Key K \<in> analz (knows Spy evs) | |
| 494 | ==> analz (Key ` (insert K H) \<union> knows Spy evs) = | |
| 495 | analz (Key ` H \<union> knows Spy evs)" | |
| 496 | by (simp add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD]) | |
| 497 | ||
| 498 | lemma analz_knows_absorb2: | |
| 499 | "Key K \<in> analz (knows Spy evs) | |
| 500 | ==> analz (Key ` (insert X (insert K H)) \<union> knows Spy evs) = | |
| 501 | analz (Key ` (insert X H) \<union> knows Spy evs)" | |
| 502 | apply (subst insert_commute) | |
| 503 | apply (erule analz_knows_absorb) | |
| 504 | done | |
| 505 | ||
| 506 | lemma analz_insert_subset_eq: | |
| 507 | "[|X \<in> analz (knows Spy evs); knows Spy evs \<subseteq> H|] | |
| 508 | ==> analz (insert X H) = analz H" | |
| 509 | apply (rule analz_insert_eq) | |
| 510 | apply (blast intro: analz_mono [THEN [2] rev_subsetD]) | |
| 511 | done | |
| 512 | ||
| 513 | lemmas analz_insert_simps = | |
| 514 | analz_insert_subset_eq Un_upper2 | |
| 515 | subset_insertI [THEN [2] subset_trans] | |
| 516 | ||
| 517 | ||
| 518 | subsection{*Freshness Lemmas*}
 | |
| 519 | ||
| 520 | lemma in_parts_Says_imp_used: | |
| 521 |      "[|Key K \<in> parts {X}; Says A B X \<in> set evs|] ==> Key K \<in> used evs"
 | |
| 522 | by (blast intro: parts_trans dest!: Says_imp_knows_Spy [THEN parts.Inj]) | |
| 523 | ||
| 524 | text{*A useful rewrite rule with @{term analz_image_keys_simps}*}
 | |
| 525 | lemma Crypt_notin_image_Key: "Crypt K X \<notin> Key ` KK" | |
| 526 | by auto | |
| 527 | ||
| 528 | lemma fresh_notin_analz_knows_Spy: | |
| 529 | "Key K \<notin> used evs ==> Key K \<notin> analz (knows Spy evs)" | |
| 530 | by (auto dest: analz_into_parts) | |
| 531 | ||
| 532 | end |