| author | wenzelm | 
| Mon, 05 Sep 2005 17:38:22 +0200 | |
| changeset 17264 | c5b280a52a67 | 
| parent 16417 | 9bc16273c2d4 | 
| child 17990 | 86d462f305e0 | 
| permissions | -rw-r--r-- | 
| 2318 | 1 | (* Title: HOL/Auth/Public | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | ||
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: 
3478diff
changeset | 6 | Theory of Public Keys (common to all public-key protocols) | 
| 2318 | 7 | |
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: 
3478diff
changeset | 8 | Private and public keys; initial states of agents | 
| 2318 | 9 | *) | 
| 10 | ||
| 16417 | 11 | theory Public imports Event begin | 
| 13922 | 12 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 13 | lemma invKey_K: "K \<in> symKeys ==> invKey K = K" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 14 | by (simp add: symKeys_def) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 15 | |
| 13922 | 16 | subsection{*Asymmetric Keys*}
 | 
| 2318 | 17 | |
| 18 | consts | |
| 13922 | 19 | (*the bool is TRUE if a signing key*) | 
| 20 | publicKey :: "[bool,agent] => key" | |
| 2318 | 21 | |
| 22 | syntax | |
| 13922 | 23 | pubEK :: "agent => key" | 
| 24 | pubSK :: "agent => key" | |
| 25 | ||
| 26 | privateKey :: "[bool,agent] => key" | |
| 27 | priEK :: "agent => key" | |
| 28 | priSK :: "agent => key" | |
| 29 | ||
| 30 | translations | |
| 31 | "pubEK" == "publicKey False" | |
| 32 | "pubSK" == "publicKey True" | |
| 33 | ||
| 34 | (*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*) | |
| 35 | "privateKey b A" == "invKey (publicKey b A)" | |
| 36 | "priEK A" == "privateKey False A" | |
| 37 | "priSK A" == "privateKey True A" | |
| 38 | ||
| 39 | ||
| 40 | text{*These translations give backward compatibility.  They represent the
 | |
| 41 | simple situation where the signature and encryption keys are the same.*} | |
| 42 | syntax | |
| 43 | pubK :: "agent => key" | |
| 44 | priK :: "agent => key" | |
| 45 | ||
| 46 | translations | |
| 47 | "pubK A" == "pubEK A" | |
| 48 | "priK A" == "invKey (pubEK A)" | |
| 49 | ||
| 50 | ||
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 51 | text{*By freeness of agents, no two agents have the same key.  Since
 | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 52 |   @{term "True\<noteq>False"}, no agent has identical signing and encryption keys*}
 | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 53 | specification (publicKey) | 
| 13922 | 54 | injective_publicKey: | 
| 55 | "publicKey b A = publicKey c A' ==> b=c & A=A'" | |
| 14261 | 56 | apply (rule exI [of _ "%b A. 2 * agent_case 0 (\<lambda>n. n + 2) 1 A + (if b then 1 else 0)"]) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 57 | apply (auto simp add: inj_on_def split: agent.split, presburger+) | 
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 58 | done | 
| 13922 | 59 | |
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 60 | |
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 61 | axioms | 
| 13922 | 62 | (*No private key equals any public key (essential to ensure that private | 
| 63 | keys are private!) *) | |
| 64 | privateKey_neq_publicKey [iff]: "privateKey b A \<noteq> publicKey c A'" | |
| 65 | ||
| 13926 | 66 | declare privateKey_neq_publicKey [THEN not_sym, iff] | 
| 13922 | 67 | |
| 68 | ||
| 13926 | 69 | subsection{*Basic properties of @{term pubK} and @{term priK}*}
 | 
| 13922 | 70 | |
| 71 | lemma [iff]: "(publicKey b A = publicKey c A') = (b=c & A=A')" | |
| 72 | by (blast dest!: injective_publicKey) | |
| 73 | ||
| 74 | lemma not_symKeys_pubK [iff]: "publicKey b A \<notin> symKeys" | |
| 13926 | 75 | by (simp add: symKeys_def) | 
| 13922 | 76 | |
| 77 | lemma not_symKeys_priK [iff]: "privateKey b A \<notin> symKeys" | |
| 13926 | 78 | by (simp add: symKeys_def) | 
| 13922 | 79 | |
| 80 | lemma symKey_neq_priEK: "K \<in> symKeys ==> K \<noteq> priEK A" | |
| 81 | by auto | |
| 82 | ||
| 83 | lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'" | |
| 13926 | 84 | by blast | 
| 13922 | 85 | |
| 86 | lemma symKeys_invKey_iff [iff]: "(invKey K \<in> symKeys) = (K \<in> symKeys)" | |
| 13926 | 87 | by (unfold symKeys_def, auto) | 
| 2318 | 88 | |
| 13922 | 89 | lemma analz_symKeys_Decrypt: | 
| 90 | "[| Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H |] | |
| 91 | ==> X \<in> analz H" | |
| 92 | by (auto simp add: symKeys_def) | |
| 93 | ||
| 94 | ||
| 95 | ||
| 96 | subsection{*"Image" equations that hold for injective functions*}
 | |
| 97 | ||
| 98 | lemma invKey_image_eq [simp]: "(invKey x \<in> invKey`A) = (x \<in> A)" | |
| 13926 | 99 | by auto | 
| 13922 | 100 | |
| 101 | (*holds because invKey is injective*) | |
| 102 | lemma publicKey_image_eq [simp]: | |
| 103 | "(publicKey b x \<in> publicKey c ` AA) = (b=c & x \<in> AA)" | |
| 104 | by auto | |
| 105 | ||
| 106 | lemma privateKey_notin_image_publicKey [simp]: "privateKey b x \<notin> publicKey c ` AA" | |
| 13926 | 107 | by auto | 
| 13922 | 108 | |
| 109 | lemma privateKey_image_eq [simp]: | |
| 110 | "(privateKey b A \<in> invKey ` publicKey c ` AS) = (b=c & A\<in>AS)" | |
| 111 | by auto | |
| 112 | ||
| 113 | lemma publicKey_notin_image_privateKey [simp]: "publicKey b A \<notin> invKey ` publicKey c ` AS" | |
| 13926 | 114 | by auto | 
| 13922 | 115 | |
| 116 | ||
| 117 | subsection{*Symmetric Keys*}
 | |
| 118 | ||
| 119 | text{*For some protocols, it is convenient to equip agents with symmetric as
 | |
| 120 | well as asymmetric keys.  The theory @{text Shared} assumes that all keys
 | |
| 121 | are symmetric.*} | |
| 122 | ||
| 123 | consts | |
| 124 |   shrK    :: "agent => key"    --{*long-term shared keys*}
 | |
| 125 | ||
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 126 | specification (shrK) | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 127 | inj_shrK: "inj shrK" | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 128 |   --{*No two agents have the same long-term key*}
 | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 129 | apply (rule exI [of _ "agent_case 0 (\<lambda>n. n + 2) 1"]) | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 130 | apply (simp add: inj_on_def split: agent.split) | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 131 | done | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 132 | |
| 13922 | 133 | axioms | 
| 134 |   sym_shrK [iff]: "shrK X \<in> symKeys" --{*All shared keys are symmetric*}
 | |
| 135 | ||
| 136 | (*Injectiveness: Agents' long-term keys are distinct.*) | |
| 137 | declare inj_shrK [THEN inj_eq, iff] | |
| 138 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 139 | lemma invKey_shrK [simp]: "invKey (shrK A) = shrK A" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 140 | by (simp add: invKey_K) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 141 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 142 | lemma analz_shrK_Decrypt: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 143 | "[| Crypt (shrK A) X \<in> analz H; Key(shrK A) \<in> analz H |] ==> X \<in> analz H" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 144 | by auto | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 145 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 146 | lemma analz_Decrypt': | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 147 | "[| Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H |] ==> X \<in> analz H" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 148 | by (auto simp add: invKey_K) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 149 | |
| 13922 | 150 | lemma priK_neq_shrK [iff]: "shrK A \<noteq> privateKey b C" | 
| 13926 | 151 | by (simp add: symKeys_neq_imp_neq) | 
| 13922 | 152 | |
| 153 | declare priK_neq_shrK [THEN not_sym, simp] | |
| 154 | ||
| 155 | lemma pubK_neq_shrK [iff]: "shrK A \<noteq> publicKey b C" | |
| 13926 | 156 | by (simp add: symKeys_neq_imp_neq) | 
| 13922 | 157 | |
| 158 | declare pubK_neq_shrK [THEN not_sym, simp] | |
| 159 | ||
| 160 | lemma priEK_noteq_shrK [simp]: "priEK A \<noteq> shrK B" | |
| 161 | by auto | |
| 162 | ||
| 163 | lemma publicKey_notin_image_shrK [simp]: "publicKey b x \<notin> shrK ` AA" | |
| 164 | by auto | |
| 165 | ||
| 166 | lemma privateKey_notin_image_shrK [simp]: "privateKey b x \<notin> shrK ` AA" | |
| 167 | by auto | |
| 168 | ||
| 169 | lemma shrK_notin_image_publicKey [simp]: "shrK x \<notin> publicKey b ` AA" | |
| 170 | by auto | |
| 171 | ||
| 172 | lemma shrK_notin_image_privateKey [simp]: "shrK x \<notin> invKey ` publicKey b ` AA" | |
| 173 | by auto | |
| 174 | ||
| 175 | lemma shrK_image_eq [simp]: "(shrK x \<in> shrK ` AA) = (x \<in> AA)" | |
| 176 | by auto | |
| 177 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 178 | text{*For some reason, moving this up can make some proofs loop!*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 179 | declare invKey_K [simp] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 180 | |
| 13922 | 181 | |
| 182 | subsection{*Initial States of Agents*}
 | |
| 183 | ||
| 184 | text{*Note: for all practical purposes, all that matters is the initial
 | |
| 185 | knowledge of the Spy. All other agents are automata, merely following the | |
| 186 | protocol.*} | |
| 2318 | 187 | |
| 5183 | 188 | primrec | 
| 2318 | 189 | (*Agents know their private key and all public keys*) | 
| 13922 | 190 | initState_Server: | 
| 191 | "initState Server = | |
| 192 |        {Key (priEK Server), Key (priSK Server)} \<union> 
 | |
| 193 | (Key ` range pubEK) \<union> (Key ` range pubSK) \<union> (Key ` range shrK)" | |
| 194 | ||
| 195 | initState_Friend: | |
| 196 | "initState (Friend i) = | |
| 197 |        {Key (priEK(Friend i)), Key (priSK(Friend i)), Key (shrK(Friend i))} \<union> 
 | |
| 198 | (Key ` range pubEK) \<union> (Key ` range pubSK)" | |
| 199 | ||
| 200 | initState_Spy: | |
| 201 | "initState Spy = | |
| 202 | (Key ` invKey ` pubEK ` bad) \<union> (Key ` invKey ` pubSK ` bad) \<union> | |
| 203 | (Key ` shrK ` bad) \<union> | |
| 204 | (Key ` range pubEK) \<union> (Key ` range pubSK)" | |
| 205 | ||
| 206 | ||
| 207 | text{*These lemmas allow reasoning about @{term "used evs"} rather than
 | |
| 13935 | 208 |    @{term "knows Spy evs"}, which is useful when there are private Notes. 
 | 
| 209 |    Because they depend upon the definition of @{term initState}, they cannot
 | |
| 210 | be moved up.*} | |
| 13922 | 211 | |
| 212 | lemma used_parts_subset_parts [rule_format]: | |
| 213 |      "\<forall>X \<in> used evs. parts {X} \<subseteq> used evs"
 | |
| 214 | apply (induct evs) | |
| 215 | prefer 2 | |
| 216 | apply (simp add: used_Cons) | |
| 217 | apply (rule ballI) | |
| 13935 | 218 | apply (case_tac a, auto) | 
| 219 | apply (auto dest!: parts_cut) | |
| 13922 | 220 | txt{*Base case*}
 | 
| 13935 | 221 | apply (simp add: used_Nil) | 
| 13922 | 222 | done | 
| 223 | ||
| 224 | lemma MPair_used_D: "{|X,Y|} \<in> used H ==> X \<in> used H & Y \<in> used H"
 | |
| 225 | by (drule used_parts_subset_parts, simp, blast) | |
| 226 | ||
| 227 | lemma MPair_used [elim!]: | |
| 228 |      "[| {|X,Y|} \<in> used H;
 | |
| 229 | [| X \<in> used H; Y \<in> used H |] ==> P |] | |
| 230 | ==> P" | |
| 231 | by (blast dest: MPair_used_D) | |
| 232 | ||
| 233 | ||
| 234 | text{*Rewrites should not refer to  @{term "initState(Friend i)"} because
 | |
| 235 | that expression is not in normal form.*} | |
| 236 | ||
| 237 | lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
 | |
| 238 | apply (unfold keysFor_def) | |
| 239 | apply (induct_tac "C") | |
| 240 | apply (auto intro: range_eqI) | |
| 241 | done | |
| 242 | ||
| 243 | lemma Crypt_notin_initState: "Crypt K X \<notin> parts (initState B)" | |
| 244 | by (induct B, auto) | |
| 245 | ||
| 13935 | 246 | lemma Crypt_notin_used_empty [simp]: "Crypt K X \<notin> used []" | 
| 247 | by (simp add: Crypt_notin_initState used_Nil) | |
| 248 | ||
| 13922 | 249 | (*** Basic properties of shrK ***) | 
| 250 | ||
| 251 | (*Agents see their own shared keys!*) | |
| 252 | lemma shrK_in_initState [iff]: "Key (shrK A) \<in> initState A" | |
| 13926 | 253 | by (induct_tac "A", auto) | 
| 13922 | 254 | |
| 255 | lemma shrK_in_knows [iff]: "Key (shrK A) \<in> knows A evs" | |
| 256 | by (simp add: initState_subset_knows [THEN subsetD]) | |
| 257 | ||
| 258 | lemma shrK_in_used [iff]: "Key (shrK A) \<in> used evs" | |
| 13926 | 259 | by (rule initState_into_used, blast) | 
| 13922 | 260 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 261 | |
| 13922 | 262 | (** Fresh keys never clash with long-term shared keys **) | 
| 263 | ||
| 264 | (*Used in parts_induct_tac and analz_Fake_tac to distinguish session keys | |
| 265 | from long-term shared keys*) | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 266 | lemma Key_not_used [simp]: "Key K \<notin> used evs ==> K \<notin> range shrK" | 
| 13926 | 267 | by blast | 
| 13922 | 268 | |
| 269 | lemma shrK_neq: "Key K \<notin> used evs ==> shrK B \<noteq> K" | |
| 13926 | 270 | by blast | 
| 13922 | 271 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 272 | declare shrK_neq [THEN not_sym, simp] | 
| 2318 | 273 | |
| 274 | ||
| 13922 | 275 | subsection{*Function @{term spies} *}
 | 
| 276 | ||
| 277 | text{*Agents see their own private keys!*}
 | |
| 278 | lemma priK_in_initState [iff]: "Key (privateKey b A) \<in> initState A" | |
| 13926 | 279 | by (induct_tac "A", auto) | 
| 13922 | 280 | |
| 281 | text{*Agents see all public keys!*}
 | |
| 282 | lemma publicKey_in_initState [iff]: "Key (publicKey b A) \<in> initState B" | |
| 13926 | 283 | by (case_tac "B", auto) | 
| 13922 | 284 | |
| 285 | text{*All public keys are visible*}
 | |
| 286 | lemma spies_pubK [iff]: "Key (publicKey b A) \<in> spies evs" | |
| 287 | apply (induct_tac "evs") | |
| 288 | apply (simp_all add: imageI knows_Cons split add: event.split) | |
| 289 | done | |
| 290 | ||
| 291 | declare spies_pubK [THEN analz.Inj, iff] | |
| 292 | ||
| 293 | text{*Spy sees private keys of bad agents!*}
 | |
| 294 | lemma Spy_spies_bad_privateKey [intro!]: | |
| 295 | "A \<in> bad ==> Key (privateKey b A) \<in> spies evs" | |
| 296 | apply (induct_tac "evs") | |
| 297 | apply (simp_all add: imageI knows_Cons split add: event.split) | |
| 298 | done | |
| 299 | ||
| 300 | text{*Spy sees long-term shared keys of bad agents!*}
 | |
| 301 | lemma Spy_spies_bad_shrK [intro!]: | |
| 302 | "A \<in> bad ==> Key (shrK A) \<in> spies evs" | |
| 303 | apply (induct_tac "evs") | |
| 304 | apply (simp_all add: imageI knows_Cons split add: event.split) | |
| 305 | done | |
| 306 | ||
| 307 | lemma publicKey_into_used [iff] :"Key (publicKey b A) \<in> used evs" | |
| 308 | apply (rule initState_into_used) | |
| 309 | apply (rule publicKey_in_initState [THEN parts.Inj]) | |
| 310 | done | |
| 311 | ||
| 312 | lemma privateKey_into_used [iff]: "Key (privateKey b A) \<in> used evs" | |
| 313 | apply(rule initState_into_used) | |
| 314 | apply(rule priK_in_initState [THEN parts.Inj]) | |
| 315 | done | |
| 316 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 317 | (*For case analysis on whether or not an agent is compromised*) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 318 | lemma Crypt_Spy_analz_bad: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 319 | "[| Crypt (shrK A) X \<in> analz (knows Spy evs); A \<in> bad |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 320 | ==> X \<in> analz (knows Spy evs)" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 321 | by force | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 322 | |
| 13922 | 323 | |
| 324 | subsection{*Fresh Nonces*}
 | |
| 325 | ||
| 326 | lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)" | |
| 13926 | 327 | by (induct_tac "B", auto) | 
| 13922 | 328 | |
| 329 | lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []" | |
| 13926 | 330 | by (simp add: used_Nil) | 
| 13922 | 331 | |
| 11104 | 332 | |
| 13922 | 333 | subsection{*Supply fresh nonces for possibility theorems*}
 | 
| 334 | ||
| 335 | text{*In any trace, there is an upper bound N on the greatest nonce in use*}
 | |
| 336 | lemma Nonce_supply_lemma: "EX N. ALL n. N<=n --> Nonce n \<notin> used evs" | |
| 337 | apply (induct_tac "evs") | |
| 13926 | 338 | apply (rule_tac x = 0 in exI) | 
| 13922 | 339 | apply (simp_all (no_asm_simp) add: used_Cons split add: event.split) | 
| 340 | apply safe | |
| 341 | apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+ | |
| 342 | done | |
| 343 | ||
| 344 | lemma Nonce_supply1: "EX N. Nonce N \<notin> used evs" | |
| 13926 | 345 | by (rule Nonce_supply_lemma [THEN exE], blast) | 
| 13922 | 346 | |
| 347 | lemma Nonce_supply: "Nonce (@ N. Nonce N \<notin> used evs) \<notin> used evs" | |
| 348 | apply (rule Nonce_supply_lemma [THEN exE]) | |
| 13926 | 349 | apply (rule someI, fast) | 
| 13922 | 350 | done | 
| 351 | ||
| 13956 | 352 | subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
 | 
| 13922 | 353 | |
| 354 | lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H"
 | |
| 13926 | 355 | by blast | 
| 11104 | 356 | |
| 13922 | 357 | lemma insert_Key_image: "insert (Key K) (Key`KK \<union> C) = Key ` (insert K KK) \<union> C" | 
| 13926 | 358 | by blast | 
| 13922 | 359 | |
| 360 | ML | |
| 361 | {*
 | |
| 362 | val Key_not_used = thm "Key_not_used"; | |
| 363 | val insert_Key_singleton = thm "insert_Key_singleton"; | |
| 364 | val insert_Key_image = thm "insert_Key_image"; | |
| 365 | *} | |
| 366 | ||
| 367 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 368 | lemma Crypt_imp_keysFor :"[|Crypt K X \<in> H; K \<in> symKeys|] ==> K \<in> keysFor H" | 
| 13926 | 369 | by (drule Crypt_imp_invKey_keysFor, simp) | 
| 13922 | 370 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 371 | text{*Lemma for the trivial direction of the if-and-only-if of the 
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 372 | Session Key Compromise Theorem*} | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 373 | lemma analz_image_freshK_lemma: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 374 | "(Key K \<in> analz (Key`nE \<union> H)) --> (K \<in> nE | Key K \<in> analz H) ==> | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 375 | (Key K \<in> analz (Key`nE \<union> H)) = (K \<in> nE | Key K \<in> analz H)" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 376 | by (blast intro: analz_mono [THEN [2] rev_subsetD]) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 377 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 378 | lemmas analz_image_freshK_simps = | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 379 |        simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 380 | disj_comms | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 381 | image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 382 | analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 383 | insert_Key_singleton | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 384 | Key_not_used insert_Key_image Un_assoc [THEN sym] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 385 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 386 | ML | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 387 | {*
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 388 | val analz_image_freshK_lemma = thm "analz_image_freshK_lemma"; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 389 | val analz_image_freshK_simps = thms "analz_image_freshK_simps"; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 390 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 391 | val analz_image_freshK_ss = | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 392 | simpset() delsimps [image_insert, image_Un] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 393 | delsimps [imp_disjL] (*reduces blow-up*) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 394 | addsimps thms "analz_image_freshK_simps" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 395 | *} | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 396 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 397 | method_setup analz_freshK = {*
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 398 | Method.no_args | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 399 | (Method.METHOD | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 400 | (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 401 | REPEAT_FIRST (rtac analz_image_freshK_lemma), | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 402 | ALLGOALS (asm_simp_tac analz_image_freshK_ss)])) *} | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 403 | "for proving the Session Key Compromise theorem" | 
| 13922 | 404 | |
| 405 | subsection{*Specialized Methods for Possibility Theorems*}
 | |
| 406 | ||
| 407 | ML | |
| 408 | {*
 | |
| 409 | val Nonce_supply = thm "Nonce_supply"; | |
| 410 | ||
| 411 | (*Tactic for possibility theorems (Isar interface)*) | |
| 412 | fun gen_possibility_tac ss state = state |> | |
| 413 | REPEAT (*omit used_Says so that Nonces start from different traces!*) | |
| 414 | (ALLGOALS (simp_tac (ss delsimps [used_Says])) | |
| 415 | THEN | |
| 416 | REPEAT_FIRST (eq_assume_tac ORELSE' | |
| 417 | resolve_tac [refl, conjI, Nonce_supply])) | |
| 418 | ||
| 419 | (*Tactic for possibility theorems (ML script version)*) | |
| 420 | fun possibility_tac state = gen_possibility_tac (simpset()) state | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 421 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 422 | (*For harder protocols (such as Recur) where we have to set up some | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 423 | nonces and keys initially*) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 424 | fun basic_possibility_tac st = st |> | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 425 | REPEAT | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 426 | (ALLGOALS (asm_simp_tac (simpset() setSolver safe_solver)) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 427 | THEN | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 428 | REPEAT_FIRST (resolve_tac [refl, conjI])) | 
| 13922 | 429 | *} | 
| 11104 | 430 | |
| 431 | method_setup possibility = {*
 | |
| 11270 
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
 paulson parents: 
11104diff
changeset | 432 | Method.ctxt_args (fn ctxt => | 
| 
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
 paulson parents: 
11104diff
changeset | 433 | Method.METHOD (fn facts => | 
| 15032 | 434 | gen_possibility_tac (local_simpset_of ctxt))) *} | 
| 11104 | 435 | "for proving possibility theorems" | 
| 2318 | 436 | |
| 437 | end |