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