| author | wenzelm | 
| Thu, 21 Jun 2007 20:07:26 +0200 | |
| changeset 23464 | bc2563c37b1a | 
| parent 21619 | dea0914773f7 | 
| child 24147 | edc90be09ac1 | 
| permissions | -rw-r--r-- | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 1 | (* Title: HOL/Auth/NSP_Bad | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | Copyright 1996 University of Cambridge | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 5 | |
| 14203 | 6 | ML{*add_path "$ISABELLE_HOME/src/HOL/Auth"*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 7 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 8 | Original file is ../Auth/NS_Public_Bad | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 9 | *) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 10 | |
| 14203 | 11 | header{*Analyzing the Needham-Schroeder Public-Key Protocol in UNITY*}
 | 
| 12 | ||
| 18556 | 13 | theory NSP_Bad imports "../Auth/Public" "../UNITY_Main" begin | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 14 | |
| 14203 | 15 | text{*This is the flawed version, vulnerable to Lowe's attack.
 | 
| 16 | From page 260 of | |
| 17 | Burrows, Abadi and Needham. A Logic of Authentication. | |
| 18 | Proc. Royal Soc. 426 (1989). | |
| 19 | *} | |
| 20 | ||
| 21 | types state = "event list" | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 22 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 23 | constdefs | 
| 14203 | 24 | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 25 | (*The spy MAY say anything he CAN say. We do not expect him to | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 26 | invent new nonces here, but he can also use NS1. Common to | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 27 | all similar protocols.*) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 28 | Fake :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 29 |     "Fake == {(s,s').
 | 
| 14203 | 30 | \<exists>B X. s' = Says Spy B X # s | 
| 31 | & X \<in> synth (analz (spies s))}" | |
| 32 | ||
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 33 | (*The numeric suffixes on A identify the rule*) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 34 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 35 | (*Alice initiates a protocol run, sending a nonce to Bob*) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 36 | NS1 :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 37 |     "NS1 == {(s1,s').
 | 
| 14203 | 38 | \<exists>A1 B NA. | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 39 | 	         s' = Says A1 B (Crypt (pubK B) {|Nonce NA, Agent A1|}) # s1
 | 
| 14203 | 40 | & Nonce NA \<notin> used s1}" | 
| 41 | ||
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 42 | (*Bob responds to Alice's message with a further nonce*) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 43 | NS2 :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 44 |     "NS2 == {(s2,s').
 | 
| 14203 | 45 | \<exists>A' A2 B NA NB. | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 46 | 	         s' = Says B A2 (Crypt (pubK A2) {|Nonce NA, Nonce NB|}) # s2
 | 
| 14203 | 47 |                & Says A' B (Crypt (pubK B) {|Nonce NA, Agent A2|}) \<in> set s2
 | 
| 48 | & Nonce NB \<notin> used s2}" | |
| 49 | ||
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 50 | (*Alice proves her existence by sending NB back to Bob.*) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 51 | NS3 :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 52 |     "NS3 == {(s3,s').
 | 
| 14203 | 53 | \<exists>A3 B' B NA NB. | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 54 | s' = Says A3 B (Crypt (pubK B) (Nonce NB)) # s3 | 
| 14203 | 55 |                & Says A3  B (Crypt (pubK B) {|Nonce NA, Agent A3|}) \<in> set s3
 | 
| 56 | 	       & Says B' A3 (Crypt (pubK A3) {|Nonce NA, Nonce NB|}) \<in> set s3}"
 | |
| 57 | ||
| 58 | ||
| 59 | constdefs | |
| 60 | Nprg :: "state program" | |
| 61 | (*Initial trace is empty*) | |
| 62 |     "Nprg == mk_total_program({[]}, {Fake, NS1, NS2, NS3}, UNIV)"
 | |
| 63 | ||
| 64 | declare spies_partsEs [elim] | |
| 65 | declare analz_into_parts [dest] | |
| 66 | declare Fake_parts_insert_in_Un [dest] | |
| 67 | ||
| 68 | text{*For other theories, e.g. Mutex and Lift, using [iff] slows proofs down.
 | |
| 69 | Here, it facilitates re-use of the Auth proofs.*} | |
| 70 | ||
| 71 | declare Fake_def [THEN def_act_simp, iff] | |
| 72 | declare NS1_def [THEN def_act_simp, iff] | |
| 73 | declare NS2_def [THEN def_act_simp, iff] | |
| 74 | declare NS3_def [THEN def_act_simp, iff] | |
| 75 | ||
| 76 | declare Nprg_def [THEN def_prg_Init, simp] | |
| 77 | ||
| 78 | ||
| 79 | text{*A "possibility property": there are traces that reach the end.
 | |
| 80 | Replace by LEADSTO proof!*} | |
| 81 | lemma "A \<noteq> B ==> | |
| 82 | \<exists>NB. \<exists>s \<in> reachable Nprg. Says A B (Crypt (pubK B) (Nonce NB)) \<in> set s" | |
| 83 | apply (intro exI bexI) | |
| 84 | apply (rule_tac [2] act = "totalize_act NS3" in reachable.Acts) | |
| 85 | apply (rule_tac [3] act = "totalize_act NS2" in reachable.Acts) | |
| 86 | apply (rule_tac [4] act = "totalize_act NS1" in reachable.Acts) | |
| 87 | apply (rule_tac [5] reachable.Init) | |
| 88 | apply (simp_all (no_asm_simp) add: Nprg_def totalize_act_def) | |
| 89 | apply auto | |
| 90 | done | |
| 91 | ||
| 92 | ||
| 93 | subsection{*Inductive Proofs about @{term ns_public}*}
 | |
| 94 | ||
| 95 | lemma ns_constrainsI: | |
| 96 |      "(!!act s s'. [| act \<in> {Id, Fake, NS1, NS2, NS3};
 | |
| 97 | (s,s') \<in> act; s \<in> A |] ==> s' \<in> A') | |
| 98 | ==> Nprg \<in> A co A'" | |
| 99 | apply (simp add: Nprg_def mk_total_program_def) | |
| 100 | apply (rule constrainsI, auto) | |
| 101 | done | |
| 102 | ||
| 103 | ||
| 104 | text{*This ML code does the inductions directly.*}
 | |
| 105 | ML{*
 | |
| 106 | val ns_constrainsI = thm "ns_constrainsI"; | |
| 21619 | 107 | val impCE = thm "impCE"; | 
| 14203 | 108 | |
| 109 | fun ns_constrains_tac(cs,ss) i = | |
| 110 | SELECT_GOAL | |
| 111 | (EVERY [REPEAT (etac Always_ConstrainsI 1), | |
| 112 | REPEAT (resolve_tac [StableI, stableI, | |
| 113 | constrains_imp_Constrains] 1), | |
| 114 | rtac ns_constrainsI 1, | |
| 115 | full_simp_tac ss 1, | |
| 116 | REPEAT (FIRSTGOAL (etac disjE)), | |
| 117 | ALLGOALS (clarify_tac (cs delrules [impI,impCE])), | |
| 118 | REPEAT (FIRSTGOAL analz_mono_contra_tac), | |
| 119 | ALLGOALS (asm_simp_tac ss)]) i; | |
| 120 | ||
| 121 | (*Tactic for proving secrecy theorems*) | |
| 122 | fun ns_induct_tac(cs,ss) = | |
| 123 | (SELECT_GOAL o EVERY) | |
| 124 | [rtac AlwaysI 1, | |
| 125 | force_tac (cs,ss) 1, | |
| 126 | (*"reachable" gets in here*) | |
| 127 | rtac (Always_reachable RS Always_ConstrainsI RS StableI) 1, | |
| 128 | ns_constrains_tac(cs,ss) 1]; | |
| 129 | *} | |
| 130 | ||
| 131 | method_setup ns_induct = {*
 | |
| 132 | Method.ctxt_args (fn ctxt => | |
| 21588 | 133 | Method.SIMPLE_METHOD' (ns_induct_tac (local_clasimpset_of ctxt))) *} | 
| 14203 | 134 | "for inductive reasoning about the Needham-Schroeder protocol" | 
| 135 | ||
| 136 | text{*Converts invariants into statements about reachable states*}
 | |
| 137 | lemmas Always_Collect_reachableD = | |
| 138 | Always_includes_reachable [THEN subsetD, THEN CollectD] | |
| 139 | ||
| 140 | text{*Spy never sees another agent's private key! (unless it's bad at start)*}
 | |
| 141 | lemma Spy_see_priK: | |
| 142 |      "Nprg \<in> Always {s. (Key (priK A) \<in> parts (spies s)) = (A \<in> bad)}"
 | |
| 143 | apply ns_induct | |
| 144 | apply blast | |
| 145 | done | |
| 146 | declare Spy_see_priK [THEN Always_Collect_reachableD, simp] | |
| 147 | ||
| 148 | lemma Spy_analz_priK: | |
| 149 |      "Nprg \<in> Always {s. (Key (priK A) \<in> analz (spies s)) = (A \<in> bad)}"
 | |
| 150 | by (rule Always_reachable [THEN Always_weaken], auto) | |
| 151 | declare Spy_analz_priK [THEN Always_Collect_reachableD, simp] | |
| 152 | ||
| 153 | ||
| 154 | subsection{*Authenticity properties obtained from NS2*}
 | |
| 155 | ||
| 156 | text{*It is impossible to re-use a nonce in both NS1 and NS2 provided the
 | |
| 157 | nonce is secret. (Honest users generate fresh nonces.)*} | |
| 158 | lemma no_nonce_NS1_NS2: | |
| 159 | "Nprg | |
| 160 |   \<in> Always {s. Crypt (pubK C) {|NA', Nonce NA|} \<in> parts (spies s) -->
 | |
| 161 |                 Crypt (pubK B) {|Nonce NA, Agent A|} \<in> parts (spies s) -->
 | |
| 162 | Nonce NA \<in> analz (spies s)}" | |
| 163 | apply ns_induct | |
| 164 | apply (blast intro: analz_insertI)+ | |
| 165 | done | |
| 166 | ||
| 167 | text{*Adding it to the claset slows down proofs...*}
 | |
| 168 | lemmas no_nonce_NS1_NS2_reachable = | |
| 169 | no_nonce_NS1_NS2 [THEN Always_Collect_reachableD, rule_format] | |
| 170 | ||
| 171 | ||
| 172 | text{*Unicity for NS1: nonce NA identifies agents A and B*}
 | |
| 173 | lemma unique_NA_lemma: | |
| 174 | "Nprg | |
| 175 |   \<in> Always {s. Nonce NA \<notin> analz (spies s) -->
 | |
| 176 |                 Crypt(pubK B) {|Nonce NA, Agent A|} \<in> parts(spies s) -->
 | |
| 177 |                 Crypt(pubK B') {|Nonce NA, Agent A'|} \<in> parts(spies s) -->
 | |
| 178 | A=A' & B=B'}" | |
| 179 | apply ns_induct | |
| 180 | apply auto | |
| 181 | txt{*Fake, NS1 are non-trivial*}
 | |
| 182 | done | |
| 183 | ||
| 184 | text{*Unicity for NS1: nonce NA identifies agents A and B*}
 | |
| 185 | lemma unique_NA: | |
| 186 |      "[| Crypt(pubK B)  {|Nonce NA, Agent A|} \<in> parts(spies s);
 | |
| 187 |          Crypt(pubK B') {|Nonce NA, Agent A'|} \<in> parts(spies s);
 | |
| 188 | Nonce NA \<notin> analz (spies s); | |
| 189 | s \<in> reachable Nprg |] | |
| 190 | ==> A=A' & B=B'" | |
| 191 | by (blast dest: unique_NA_lemma [THEN Always_Collect_reachableD]) | |
| 192 | ||
| 193 | ||
| 194 | text{*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure*}
 | |
| 195 | lemma Spy_not_see_NA: | |
| 196 | "[| A \<notin> bad; B \<notin> bad |] | |
| 197 | ==> Nprg \<in> Always | |
| 198 |               {s. Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) \<in> set s
 | |
| 199 | --> Nonce NA \<notin> analz (spies s)}" | |
| 200 | apply ns_induct | |
| 201 | txt{*NS3*}
 | |
| 202 | prefer 4 apply (blast intro: no_nonce_NS1_NS2_reachable) | |
| 203 | txt{*NS2*}
 | |
| 204 | prefer 3 apply (blast dest: unique_NA) | |
| 205 | txt{*NS1*}
 | |
| 206 | prefer 2 apply blast | |
| 207 | txt{*Fake*}
 | |
| 208 | apply spy_analz | |
| 209 | done | |
| 210 | ||
| 211 | ||
| 212 | text{*Authentication for A: if she receives message 2 and has used NA
 | |
| 213 | to start a run, then B has sent message 2.*} | |
| 214 | lemma A_trusts_NS2: | |
| 215 | "[| A \<notin> bad; B \<notin> bad |] | |
| 216 | ==> Nprg \<in> Always | |
| 217 |               {s. Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) \<in> set s &
 | |
| 218 |                   Crypt(pubK A) {|Nonce NA, Nonce NB|} \<in> parts (knows Spy s)
 | |
| 219 |          --> Says B A (Crypt(pubK A) {|Nonce NA, Nonce NB|}) \<in> set s}"
 | |
| 220 | (*insert an invariant for use in some of the subgoals*) | |
| 221 | apply (insert Spy_not_see_NA [of A B NA], simp, ns_induct) | |
| 222 | apply (auto dest: unique_NA) | |
| 223 | done | |
| 224 | ||
| 225 | ||
| 226 | text{*If the encrypted message appears then it originated with Alice in NS1*}
 | |
| 227 | lemma B_trusts_NS1: | |
| 228 | "Nprg \<in> Always | |
| 229 |               {s. Nonce NA \<notin> analz (spies s) -->
 | |
| 230 |                   Crypt (pubK B) {|Nonce NA, Agent A|} \<in> parts (spies s)
 | |
| 231 |          --> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) \<in> set s}"
 | |
| 232 | apply ns_induct | |
| 233 | apply blast | |
| 234 | done | |
| 235 | ||
| 236 | ||
| 237 | subsection{*Authenticity properties obtained from NS2*}
 | |
| 238 | ||
| 239 | text{*Unicity for NS2: nonce NB identifies nonce NA and agent A.
 | |
| 240 |    Proof closely follows that of @{text unique_NA}.*}
 | |
| 241 | lemma unique_NB_lemma: | |
| 242 | "Nprg | |
| 243 |   \<in> Always {s. Nonce NB \<notin> analz (spies s)  -->
 | |
| 244 |                 Crypt (pubK A) {|Nonce NA, Nonce NB|} \<in> parts (spies s) -->
 | |
| 245 |                 Crypt(pubK A'){|Nonce NA', Nonce NB|} \<in> parts(spies s) -->
 | |
| 246 | A=A' & NA=NA'}" | |
| 247 | apply ns_induct | |
| 248 | apply auto | |
| 249 | txt{*Fake, NS2 are non-trivial*}
 | |
| 250 | done | |
| 251 | ||
| 252 | lemma unique_NB: | |
| 253 |      "[| Crypt(pubK A) {|Nonce NA, Nonce NB|} \<in> parts(spies s);
 | |
| 254 |          Crypt(pubK A'){|Nonce NA', Nonce NB|} \<in> parts(spies s);
 | |
| 255 | Nonce NB \<notin> analz (spies s); | |
| 256 | s \<in> reachable Nprg |] | |
| 257 | ==> A=A' & NA=NA'" | |
| 258 | apply (blast dest: unique_NB_lemma [THEN Always_Collect_reachableD]) | |
| 259 | done | |
| 260 | ||
| 261 | ||
| 262 | text{*NB remains secret PROVIDED Alice never responds with round 3*}
 | |
| 263 | lemma Spy_not_see_NB: | |
| 264 | "[| A \<notin> bad; B \<notin> bad |] | |
| 265 | ==> Nprg \<in> Always | |
| 266 |               {s. Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s &
 | |
| 267 | (ALL C. Says A C (Crypt (pubK C) (Nonce NB)) \<notin> set s) | |
| 268 | --> Nonce NB \<notin> analz (spies s)}" | |
| 269 | apply ns_induct | |
| 270 | apply (simp_all (no_asm_simp) add: all_conj_distrib) | |
| 271 | txt{*NS3: because NB determines A*}
 | |
| 272 | prefer 4 apply (blast dest: unique_NB) | |
| 273 | txt{*NS2: by freshness and unicity of NB*}
 | |
| 274 | prefer 3 apply (blast intro: no_nonce_NS1_NS2_reachable) | |
| 275 | txt{*NS1: by freshness*}
 | |
| 276 | prefer 2 apply blast | |
| 277 | txt{*Fake*}
 | |
| 278 | apply spy_analz | |
| 279 | done | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 280 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 281 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 282 | |
| 14203 | 283 | text{*Authentication for B: if he receives message 3 and has used NB
 | 
| 284 | in message 2, then A has sent message 3--to somebody....*} | |
| 285 | lemma B_trusts_NS3: | |
| 286 | "[| A \<notin> bad; B \<notin> bad |] | |
| 287 | ==> Nprg \<in> Always | |
| 288 |               {s. Crypt (pubK B) (Nonce NB) \<in> parts (spies s) &
 | |
| 289 |                   Says B A  (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s
 | |
| 290 | --> (\<exists>C. Says A C (Crypt (pubK C) (Nonce NB)) \<in> set s)}" | |
| 291 | (*insert an invariant for use in some of the subgoals*) | |
| 292 | apply (insert Spy_not_see_NB [of A B NA NB], simp, ns_induct) | |
| 293 | apply (simp_all (no_asm_simp) add: ex_disj_distrib) | |
| 294 | apply auto | |
| 295 | txt{*NS3: because NB determines A. This use of @{text unique_NB} is robust.*}
 | |
| 296 | apply (blast intro: unique_NB [THEN conjunct1]) | |
| 297 | done | |
| 298 | ||
| 299 | ||
| 300 | text{*Can we strengthen the secrecy theorem?  NO*}
 | |
| 301 | lemma "[| A \<notin> bad; B \<notin> bad |] | |
| 302 | ==> Nprg \<in> Always | |
| 303 |               {s. Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s
 | |
| 304 | --> Nonce NB \<notin> analz (spies s)}" | |
| 305 | apply ns_induct | |
| 306 | apply auto | |
| 307 | txt{*Fake*}
 | |
| 308 | apply spy_analz | |
| 309 | txt{*NS2: by freshness and unicity of NB*}
 | |
| 310 | apply (blast intro: no_nonce_NS1_NS2_reachable) | |
| 311 | txt{*NS3: unicity of NB identifies A and NA, but not B*}
 | |
| 312 | apply (frule_tac A'=A in Says_imp_spies [THEN parts.Inj, THEN unique_NB]) | |
| 313 | apply (erule Says_imp_spies [THEN parts.Inj], auto) | |
| 314 | apply (rename_tac s B' C) | |
| 315 | txt{*This is the attack!
 | |
| 316 | @{subgoals[display,indent=0,margin=65]}
 | |
| 317 | *} | |
| 318 | oops | |
| 319 | ||
| 320 | ||
| 321 | (* | |
| 322 | THIS IS THE ATTACK! | |
| 323 | [| A \<notin> bad; B \<notin> bad |] | |
| 324 | ==> Nprg | |
| 325 | \<in> Always | |
| 326 |        {s. Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s -->
 | |
| 327 | Nonce NB \<notin> analz (knows Spy s)} | |
| 328 | 1. !!s B' C. | |
| 329 | [| A \<notin> bad; B \<notin> bad; s \<in> reachable Nprg | |
| 330 |           Says A C (Crypt (pubK C) {|Nonce NA, Agent A|}) \<in> set s;
 | |
| 331 |           Says B' A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s;
 | |
| 332 |           C \<in> bad; Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s;
 | |
| 333 | Nonce NB \<notin> analz (knows Spy s) |] | |
| 334 | ==> False | |
| 335 | *) | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 336 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 337 | end |