| author | wenzelm | 
| Mon, 12 Jun 2006 20:58:25 +0200 | |
| changeset 19859 | e5c12b5cb940 | 
| parent 18886 | 9f27383426db | 
| child 20768 | 1d478c2d621f | 
| permissions | -rw-r--r-- | 
| 5053 | 1 | (* Title: HOL/Auth/Kerberos_BAN | 
| 2 | ID: $Id$ | |
| 3 | Author: Giampaolo Bella, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 7 | header{*The Kerberos Protocol, BAN Version*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 8 | |
| 16417 | 9 | theory Kerberos_BAN imports Public begin | 
| 5053 | 10 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 11 | text{*From page 251 of
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 12 | Burrows, Abadi and Needham (1989). A Logic of Authentication. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 13 | Proc. Royal Soc. 426 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 14 | |
| 18886 | 15 | Confidentiality (secrecy) and authentication properties are also | 
| 16 | given in a termporal version: strong guarantees in a little abstracted | |
| 17 | - but very realistic - model. | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 18 | *} | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 19 | |
| 18886 | 20 | (* Temporal model of accidents: session keys can be leaked | 
| 21 | ONLY when they have expired *) | |
| 5053 | 22 | |
| 23 | syntax | |
| 13926 | 24 | CT :: "event list=>nat" | 
| 18886 | 25 | expiredK :: "[nat, event list] => bool" | 
| 26 | expiredA :: "[nat, event list] => bool" | |
| 5053 | 27 | |
| 28 | consts | |
| 29 | ||
| 30 | (*Duration of the session key*) | |
| 18886 | 31 | sesKlife :: nat | 
| 5053 | 32 | |
| 33 | (*Duration of the authenticator*) | |
| 18886 | 34 | authlife :: nat | 
| 5053 | 35 | |
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13926diff
changeset | 36 | text{*The ticket should remain fresh for two journeys on the network at least*}
 | 
| 18886 | 37 | specification (sesKlife) | 
| 38 | sesKlife_LB [iff]: "2 \<le> sesKlife" | |
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13926diff
changeset | 39 | by blast | 
| 5053 | 40 | |
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13926diff
changeset | 41 | text{*The authenticator only for one journey*}
 | 
| 18886 | 42 | specification (authlife) | 
| 43 | authlife_LB [iff]: "0 < authlife" | |
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13926diff
changeset | 44 | by blast | 
| 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13926diff
changeset | 45 | |
| 5053 | 46 | |
| 47 | translations | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 48 | "CT" == "length " | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 49 | |
| 18886 | 50 | "expiredK T evs" == "sesKlife + T < CT evs" | 
| 51 | ||
| 52 | "expiredA T evs" == "authlife + T < CT evs" | |
| 53 | ||
| 54 | ||
| 55 | constdefs | |
| 5053 | 56 | |
| 18886 | 57 | (* A is the true creator of X if she has sent X and X never appeared on | 
| 58 | the trace before this event. Recall that traces grow from head. *) | |
| 59 | Issues :: "[agent, agent, msg, event list] => bool" | |
| 60 |              ("_ Issues _ with _ on _")
 | |
| 61 | "A Issues B with X on evs == | |
| 62 |       \<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
 | |
| 63 | X \<notin> parts (spies (takeWhile (% z. z \<noteq> Says A B Y) (rev evs)))" | |
| 5053 | 64 | |
| 18886 | 65 | (* Yields the subtrace of a given trace from its beginning to a given event *) | 
| 66 |   before :: "[event, event list] => event list" ("before _ on _")
 | |
| 67 | "before ev on evs == takeWhile (% z. z ~= ev) (rev evs)" | |
| 68 | ||
| 69 | (* States than an event really appears only once on a trace *) | |
| 70 |   Unique :: "[event, event list] => bool" ("Unique _ on _")
 | |
| 71 | "Unique ev on evs == | |
| 72 | ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))" | |
| 73 | ||
| 74 | ||
| 75 | consts bankerberos :: "event list set" | |
| 76 | inductive "bankerberos" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 77 | intros | 
| 13926 | 78 | |
| 18886 | 79 | Nil: "[] \<in> bankerberos" | 
| 13926 | 80 | |
| 18886 | 81 | Fake: "\<lbrakk> evsf \<in> bankerberos; X \<in> synth (analz (spies evsf)) \<rbrakk> | 
| 82 | \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos" | |
| 13926 | 83 | |
| 84 | ||
| 18886 | 85 | BK1: "\<lbrakk> evs1 \<in> bankerberos \<rbrakk> | 
| 86 | \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1 | |
| 87 | \<in> bankerberos" | |
| 88 | ||
| 89 | ||
| 90 | BK2: "\<lbrakk> evs2 \<in> bankerberos; Key K \<notin> used evs2; K \<in> symKeys; | |
| 91 | Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk> | |
| 92 | \<Longrightarrow> Says Server A | |
| 13926 | 93 | (Crypt (shrK A) | 
| 18886 | 94 | \<lbrace>Number (CT evs2), Agent B, Key K, | 
| 95 | (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>) | |
| 96 | # evs2 \<in> bankerberos" | |
| 13926 | 97 | |
| 98 | ||
| 18886 | 99 | BK3: "\<lbrakk> evs3 \<in> bankerberos; | 
| 100 | Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) | |
| 13926 | 101 | \<in> set evs3; | 
| 18886 | 102 | Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3; | 
| 103 | \<not> expiredK Tk evs3 \<rbrakk> | |
| 104 | \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace> | |
| 105 | # evs3 \<in> bankerberos" | |
| 13926 | 106 | |
| 107 | ||
| 18886 | 108 | BK4: "\<lbrakk> evs4 \<in> bankerberos; | 
| 109 | Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>), | |
| 110 | (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4; | |
| 111 | \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk> | |
| 112 | \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4 | |
| 113 | \<in> bankerberos" | |
| 5053 | 114 | |
| 13926 | 115 | (*Old session keys may become compromised*) | 
| 18886 | 116 | Oops: "\<lbrakk> evso \<in> bankerberos; | 
| 117 | Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) | |
| 13926 | 118 | \<in> set evso; | 
| 18886 | 119 | expiredK Tk evso \<rbrakk> | 
| 120 | \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos" | |
| 13926 | 121 | |
| 122 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 123 | declare Says_imp_knows_Spy [THEN parts.Inj, dest] | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 124 | declare parts.Body [dest] | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 125 | declare analz_into_parts [dest] | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 126 | declare Fake_parts_insert_in_Un [dest] | 
| 13926 | 127 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 128 | text{*A "possibility property": there are traces that reach the end.*}
 | 
| 18886 | 129 | lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk> | 
| 130 | \<Longrightarrow> \<exists>Timestamp. \<exists>evs \<in> bankerberos. | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 131 | Says B A (Crypt K (Number Timestamp)) | 
| 13926 | 132 | \<in> set evs" | 
| 18886 | 133 | apply (cut_tac sesKlife_LB) | 
| 13926 | 134 | apply (intro exI bexI) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 135 | apply (rule_tac [2] | 
| 18886 | 136 | bankerberos.Nil [THEN bankerberos.BK1, THEN bankerberos.BK2, | 
| 137 | THEN bankerberos.BK3, THEN bankerberos.BK4]) | |
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 138 | apply (possibility, simp_all (no_asm_simp) add: used_Cons) | 
| 13926 | 139 | done | 
| 140 | ||
| 18886 | 141 | subsection{*Lemmas for reasoning about predicate "Issues"*}
 | 
| 13926 | 142 | |
| 18886 | 143 | lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)" | 
| 144 | apply (induct_tac "evs") | |
| 145 | apply (induct_tac [2] "a", auto) | |
| 146 | done | |
| 147 | ||
| 148 | lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs" | |
| 149 | apply (induct_tac "evs") | |
| 150 | apply (induct_tac [2] "a", auto) | |
| 151 | done | |
| 152 | ||
| 153 | lemma spies_Notes_rev: "spies (evs @ [Notes A X]) = | |
| 154 | (if A:bad then insert X (spies evs) else spies evs)" | |
| 155 | apply (induct_tac "evs") | |
| 156 | apply (induct_tac [2] "a", auto) | |
| 157 | done | |
| 158 | ||
| 159 | lemma spies_evs_rev: "spies evs = spies (rev evs)" | |
| 160 | apply (induct_tac "evs") | |
| 161 | apply (induct_tac [2] "a") | |
| 162 | apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev) | |
| 163 | done | |
| 164 | ||
| 165 | lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono] | |
| 166 | ||
| 167 | lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs" | |
| 168 | apply (induct_tac "evs") | |
| 169 | apply (induct_tac [2] "a", auto) | |
| 170 | txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 | |
| 171 | done | |
| 172 | ||
| 173 | lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono] | |
| 174 | ||
| 175 | ||
| 176 | text{*Lemmas for reasoning about predicate "before"*}
 | |
| 177 | lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)";
 | |
| 178 | apply (induct_tac "evs") | |
| 179 | apply simp | |
| 180 | apply (induct_tac "a") | |
| 181 | apply auto | |
| 182 | done | |
| 13926 | 183 | |
| 18886 | 184 | lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)";
 | 
| 185 | apply (induct_tac "evs") | |
| 186 | apply simp | |
| 187 | apply (induct_tac "a") | |
| 188 | apply auto | |
| 189 | done | |
| 190 | ||
| 191 | lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs"; | |
| 192 | apply (induct_tac "evs") | |
| 193 | apply simp | |
| 194 | apply (induct_tac "a") | |
| 195 | apply auto | |
| 196 | done | |
| 197 | ||
| 198 | lemma used_evs_rev: "used evs = used (rev evs)" | |
| 199 | apply (induct_tac "evs") | |
| 200 | apply simp | |
| 201 | apply (induct_tac "a") | |
| 202 | apply (simp add: used_Says_rev) | |
| 203 | apply (simp add: used_Gets_rev) | |
| 204 | apply (simp add: used_Notes_rev) | |
| 205 | done | |
| 206 | ||
| 207 | lemma used_takeWhile_used [rule_format]: | |
| 208 | "x : used (takeWhile P X) --> x : used X" | |
| 209 | apply (induct_tac "X") | |
| 210 | apply simp | |
| 211 | apply (induct_tac "a") | |
| 212 | apply (simp_all add: used_Nil) | |
| 213 | apply (blast dest!: initState_into_used)+ | |
| 214 | done | |
| 215 | ||
| 216 | lemma set_evs_rev: "set evs = set (rev evs)" | |
| 217 | apply auto | |
| 218 | done | |
| 219 | ||
| 220 | lemma takeWhile_void [rule_format]: | |
| 221 | "x \<notin> set evs \<longrightarrow> takeWhile (\<lambda>z. z \<noteq> x) evs = evs" | |
| 222 | apply auto | |
| 223 | done | |
| 224 | ||
| 225 | (**** Inductive proofs about bankerberos ****) | |
| 226 | ||
| 227 | text{*Forwarding Lemma for reasoning about the encrypted portion of message BK3*}
 | |
| 228 | lemma BK3_msg_in_parts_spies: | |
| 229 | "Says S A (Crypt KA \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs | |
| 230 | \<Longrightarrow> X \<in> parts (spies evs)" | |
| 231 | apply blast | |
| 232 | done | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 233 | |
| 13926 | 234 | lemma Oops_parts_spies: | 
| 18886 | 235 | "Says Server A (Crypt (shrK A) \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs | 
| 236 | \<Longrightarrow> K \<in> parts (spies evs)" | |
| 237 | apply blast | |
| 238 | done | |
| 13926 | 239 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 240 | text{*Spy never sees another agent's shared key! (unless it's bad at start)*}
 | 
| 13926 | 241 | lemma Spy_see_shrK [simp]: | 
| 18886 | 242 | "evs \<in> bankerberos \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" | 
| 243 | apply (erule bankerberos.induct) | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 244 | apply (frule_tac [7] Oops_parts_spies) | 
| 18886 | 245 | apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast+) | 
| 13926 | 246 | done | 
| 5053 | 247 | |
| 248 | ||
| 13926 | 249 | lemma Spy_analz_shrK [simp]: | 
| 18886 | 250 | "evs \<in> bankerberos \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)" | 
| 251 | apply auto | |
| 252 | done | |
| 13926 | 253 | |
| 254 | lemma Spy_see_shrK_D [dest!]: | |
| 18886 | 255 | "\<lbrakk> Key (shrK A) \<in> parts (spies evs); | 
| 256 | evs \<in> bankerberos \<rbrakk> \<Longrightarrow> A:bad" | |
| 257 | apply (blast dest: Spy_see_shrK) | |
| 258 | done | |
| 13926 | 259 | |
| 260 | lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!] | |
| 261 | ||
| 262 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 263 | text{*Nobody can have used non-existent keys!*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 264 | lemma new_keys_not_used [simp]: | 
| 18886 | 265 | "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> bankerberos\<rbrakk> | 
| 266 | \<Longrightarrow> K \<notin> keysFor (parts (spies evs))" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 267 | apply (erule rev_mp) | 
| 18886 | 268 | apply (erule bankerberos.induct) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 269 | apply (frule_tac [7] Oops_parts_spies) | 
| 18886 | 270 | apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 271 | txt{*Fake*}
 | 
| 13926 | 272 | apply (force dest!: keysFor_parts_insert) | 
| 18886 | 273 | txt{*BK2, BK3, BK4*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 274 | apply (force dest!: analz_shrK_Decrypt)+ | 
| 13926 | 275 | done | 
| 276 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 277 | subsection{* Lemmas concerning the form of items passed in messages *}
 | 
| 13926 | 278 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 279 | text{*Describes the form of K, X and K' when the Server sends this message.*}
 | 
| 13926 | 280 | lemma Says_Server_message_form: | 
| 18886 | 281 | "\<lbrakk> Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) | 
| 282 | \<in> set evs; evs \<in> bankerberos \<rbrakk> | |
| 283 | \<Longrightarrow> K' = shrK A & K \<notin> range shrK & | |
| 284 | Ticket = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>) & | |
| 285 | Key K \<notin> used(before | |
| 286 | Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) | |
| 287 | on evs) & | |
| 288 | Tk = CT(before | |
| 289 | Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) | |
| 290 | on evs)" | |
| 291 | apply (unfold before_def) | |
| 13926 | 292 | apply (erule rev_mp) | 
| 18886 | 293 | apply (erule bankerberos.induct, simp_all) | 
| 294 | txt{*We need this simplification only for Message 2*}
 | |
| 295 | apply (simp (no_asm) add: takeWhile_tail) | |
| 296 | apply auto | |
| 297 | txt{*Two subcases of Message 2. Subcase: used before*}
 | |
| 298 | apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] | |
| 299 | used_takeWhile_used) | |
| 300 | txt{*subcase: CT before*}
 | |
| 301 | apply (fastsimp dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void]) | |
| 13926 | 302 | done | 
| 5053 | 303 | |
| 304 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 305 | text{*If the encrypted message appears then it originated with the Server
 | 
| 13926 | 306 | PROVIDED that A is NOT compromised! | 
| 18886 | 307 | This allows A to verify freshness of the session key. | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 308 | *} | 
| 18886 | 309 | lemma Kab_authentic: | 
| 310 | "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 311 | \<in> parts (spies evs); | 
| 18886 | 312 | A \<notin> bad; evs \<in> bankerberos \<rbrakk> | 
| 313 | \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) | |
| 13926 | 314 | \<in> set evs" | 
| 315 | apply (erule rev_mp) | |
| 18886 | 316 | apply (erule bankerberos.induct) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 317 | apply (frule_tac [7] Oops_parts_spies) | 
| 18886 | 318 | apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast) | 
| 13926 | 319 | done | 
| 320 | ||
| 321 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 322 | text{*If the TICKET appears then it originated with the Server*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 323 | text{*FRESHNESS OF THE SESSION KEY to B*}
 | 
| 18886 | 324 | lemma ticket_authentic: | 
| 325 | "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs); | |
| 326 | B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 327 | \<Longrightarrow> Says Server A | |
| 328 | (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, | |
| 329 | Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>) | |
| 13926 | 330 | \<in> set evs" | 
| 331 | apply (erule rev_mp) | |
| 18886 | 332 | apply (erule bankerberos.induct) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 333 | apply (frule_tac [7] Oops_parts_spies) | 
| 18886 | 334 | apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast) | 
| 13926 | 335 | done | 
| 336 | ||
| 337 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 338 | text{*EITHER describes the form of X when the following message is sent,
 | 
| 13926 | 339 | OR reduces it to the Fake case. | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 340 |   Use @{text Says_Server_message_form} if applicable.*}
 | 
| 13926 | 341 | lemma Says_S_message_form: | 
| 18886 | 342 | "\<lbrakk> Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 343 | \<in> set evs; | 
| 18886 | 344 | evs \<in> bankerberos \<rbrakk> | 
| 345 | \<Longrightarrow> (K \<notin> range shrK & X = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>)) | |
| 13926 | 346 | | X \<in> analz (spies evs)" | 
| 347 | apply (case_tac "A \<in> bad") | |
| 348 | apply (force dest!: Says_imp_spies [THEN analz.Inj]) | |
| 349 | apply (frule Says_imp_spies [THEN parts.Inj]) | |
| 18886 | 350 | apply (blast dest!: Kab_authentic Says_Server_message_form) | 
| 13926 | 351 | done | 
| 352 | ||
| 5053 | 353 | |
| 354 | ||
| 13926 | 355 | (**** | 
| 356 | The following is to prove theorems of the form | |
| 357 | ||
| 18886 | 358 | Key K \<in> analz (insert (Key KAB) (spies evs)) \<Longrightarrow> | 
| 13926 | 359 | Key K \<in> analz (spies evs) | 
| 360 | ||
| 361 | A more general formula must be proved inductively. | |
| 362 | ||
| 363 | ****) | |
| 364 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 365 | text{* Session keys are not used to encrypt other session keys *}
 | 
| 13926 | 366 | lemma analz_image_freshK [rule_format (no_asm)]: | 
| 18886 | 367 | "evs \<in> bankerberos \<Longrightarrow> | 
| 368 | \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow> | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 369 | (Key K \<in> analz (Key`KK Un (spies evs))) = | 
| 13926 | 370 | (K \<in> KK | Key K \<in> analz (spies evs))" | 
| 18886 | 371 | apply (erule bankerberos.induct) | 
| 13926 | 372 | apply (drule_tac [7] Says_Server_message_form) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 373 | apply (erule_tac [5] Says_S_message_form [THEN disjE], analz_freshK, spy_analz, auto) | 
| 13926 | 374 | done | 
| 375 | ||
| 376 | ||
| 377 | lemma analz_insert_freshK: | |
| 18886 | 378 | "\<lbrakk> evs \<in> bankerberos; KAB \<notin> range shrK \<rbrakk> \<Longrightarrow> | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 379 | (Key K \<in> analz (insert (Key KAB) (spies evs))) = | 
| 13926 | 380 | (K = KAB | Key K \<in> analz (spies evs))" | 
| 18886 | 381 | apply (simp only: analz_image_freshK analz_image_freshK_simps) | 
| 382 | done | |
| 13926 | 383 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 384 | text{* The session key K uniquely identifies the message *}
 | 
| 13926 | 385 | lemma unique_session_keys: | 
| 18886 | 386 | "\<lbrakk> Says Server A | 
| 387 | (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs; | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 388 | Says Server A' | 
| 18886 | 389 | (Crypt (shrK A') \<lbrace>Number Tk', Agent B', Key K, X'\<rbrace>) \<in> set evs; | 
| 390 | evs \<in> bankerberos \<rbrakk> \<Longrightarrow> A=A' & Tk=Tk' & B=B' & X = X'" | |
| 13926 | 391 | apply (erule rev_mp) | 
| 392 | apply (erule rev_mp) | |
| 18886 | 393 | apply (erule bankerberos.induct) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 394 | apply (frule_tac [7] Oops_parts_spies) | 
| 18886 | 395 | apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all) | 
| 396 | txt{*BK2: it can't be a new key*}
 | |
| 397 | apply blast | |
| 398 | done | |
| 399 | ||
| 400 | lemma Server_Unique: | |
| 401 | "\<lbrakk> Says Server A | |
| 402 | (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs; | |
| 403 | evs \<in> bankerberos \<rbrakk> \<Longrightarrow> | |
| 404 | Unique Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) | |
| 405 | on evs" | |
| 406 | apply (erule rev_mp, erule bankerberos.induct, simp_all add: Unique_def) | |
| 13926 | 407 | apply blast | 
| 408 | done | |
| 409 | ||
| 410 | ||
| 18886 | 411 | subsection{*Non-temporal guarantees, explicitly relying on non-occurrence of
 | 
| 412 | oops events - refined below by temporal guarantees*} | |
| 413 | ||
| 414 | text{*Non temporal treatment of confidentiality*}
 | |
| 415 | ||
| 416 | text{* Lemma: the session key sent in msg BK2 would be lost by oops
 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 417 | if the spy could see it! *} | 
| 18886 | 418 | lemma lemma_conf [rule_format (no_asm)]: | 
| 419 | "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 420 | \<Longrightarrow> Says Server A | |
| 421 | (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, | |
| 422 | Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>) | |
| 423 | \<in> set evs \<longrightarrow> | |
| 424 | Key K \<in> analz (spies evs) \<longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<in> set evs" | |
| 425 | apply (erule bankerberos.induct) | |
| 426 | apply (frule_tac [7] Says_Server_message_form) | |
| 427 | apply (frule_tac [5] Says_S_message_form [THEN disjE]) | |
| 428 | apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes) | |
| 429 | txt{*Fake*}
 | |
| 430 | apply spy_analz | |
| 431 | txt{*BK2*}
 | |
| 432 | apply (blast intro: parts_insertI) | |
| 433 | txt{*BK3*}
 | |
| 434 | apply (case_tac "Aa \<in> bad") | |
| 435 | prefer 2 apply (blast dest: Kab_authentic unique_session_keys) | |
| 436 | apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz) | |
| 437 | txt{*Oops*}
 | |
| 438 | apply (blast dest: unique_session_keys) | |
| 439 | done | |
| 440 | ||
| 441 | ||
| 442 | text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
 | |
| 443 | as long as they have not expired.*} | |
| 444 | lemma Confidentiality_S: | |
| 445 | "\<lbrakk> Says Server A | |
| 446 | (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs; | |
| 447 | Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs; | |
| 448 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos | |
| 449 | \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)" | |
| 450 | apply (frule Says_Server_message_form, assumption) | |
| 451 | apply (blast intro: lemma_conf) | |
| 452 | done | |
| 453 | ||
| 454 | text{*Confidentiality for Alice*}
 | |
| 455 | lemma Confidentiality_A: | |
| 456 | "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | |
| 457 | Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs; | |
| 458 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos | |
| 459 | \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)" | |
| 460 | apply (blast dest!: Kab_authentic Confidentiality_S) | |
| 461 | done | |
| 462 | ||
| 463 | text{*Confidentiality for Bob*}
 | |
| 464 | lemma Confidentiality_B: | |
| 465 | "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> | |
| 466 | \<in> parts (spies evs); | |
| 467 | Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs; | |
| 468 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos | |
| 469 | \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)" | |
| 470 | apply (blast dest!: ticket_authentic Confidentiality_S) | |
| 471 | done | |
| 472 | ||
| 473 | text{*Non temporal treatment of authentication*}
 | |
| 13926 | 474 | |
| 18886 | 475 | text{*Lemmas @{text lemma_A} and @{text lemma_B} in fact are common to both temporal and non-temporal treatments*}
 | 
| 476 | lemma lemma_A [rule_format]: | |
| 477 | "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 478 | \<Longrightarrow> | |
| 479 | Key K \<notin> analz (spies evs) \<longrightarrow> | |
| 480 | Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) | |
| 481 | \<in> set evs \<longrightarrow> | |
| 482 | Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs) \<longrightarrow> | |
| 483 | Says A B \<lbrace>X, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> | |
| 484 | \<in> set evs" | |
| 485 | apply (erule bankerberos.induct) | |
| 486 | apply (frule_tac [7] Oops_parts_spies) | |
| 487 | apply (frule_tac [5] Says_S_message_form) | |
| 488 | apply (frule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra) | |
| 489 | apply (simp_all (no_asm_simp) add: all_conj_distrib) | |
| 490 | txt{*Fake*}
 | |
| 491 | apply blast | |
| 492 | txt{*BK2*}
 | |
| 493 | apply (force dest: Crypt_imp_invKey_keysFor) | |
| 494 | txt{*BK3*}
 | |
| 495 | apply (blast dest: Kab_authentic unique_session_keys) | |
| 496 | done | |
| 497 | lemma lemma_B [rule_format]: | |
| 498 | "\<lbrakk> B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 499 | \<Longrightarrow> Key K \<notin> analz (spies evs) \<longrightarrow> | |
| 500 | Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) | |
| 501 | \<in> set evs \<longrightarrow> | |
| 502 | Crypt K (Number Ta) \<in> parts (spies evs) \<longrightarrow> | |
| 503 | Says B A (Crypt K (Number Ta)) \<in> set evs" | |
| 504 | apply (erule bankerberos.induct) | |
| 505 | apply (frule_tac [7] Oops_parts_spies) | |
| 506 | apply (frule_tac [5] Says_S_message_form) | |
| 507 | apply (drule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra) | |
| 508 | apply (simp_all (no_asm_simp) add: all_conj_distrib) | |
| 509 | txt{*Fake*}
 | |
| 510 | apply blast | |
| 511 | txt{*BK2*} 
 | |
| 512 | apply (force dest: Crypt_imp_invKey_keysFor) | |
| 513 | txt{*BK4*}
 | |
| 514 | apply (blast dest: ticket_authentic unique_session_keys | |
| 515 | Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad) | |
| 516 | done | |
| 517 | ||
| 518 | ||
| 519 | text{*The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
 | |
| 520 | ||
| 521 | ||
| 522 | text{*Authentication of A to B*}
 | |
| 523 | lemma B_authenticates_A_r: | |
| 524 | "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs); | |
| 525 | Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs); | |
| 526 | Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs; | |
| 527 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 528 | \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>, | |
| 529 | Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs" | |
| 530 | apply (blast dest!: ticket_authentic | |
| 531 | intro!: lemma_A | |
| 532 | elim!: Confidentiality_S [THEN [2] rev_notE]) | |
| 533 | done | |
| 534 | ||
| 535 | ||
| 536 | text{*Authentication of B to A*}
 | |
| 537 | lemma A_authenticates_B_r: | |
| 538 | "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs); | |
| 539 | Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | |
| 540 | Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs; | |
| 541 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 542 | \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs" | |
| 543 | apply (blast dest!: Kab_authentic | |
| 544 | intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE]) | |
| 545 | done | |
| 546 | ||
| 547 | lemma B_authenticates_A: | |
| 548 | "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs); | |
| 549 | Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs); | |
| 550 | Key K \<notin> analz (spies evs); | |
| 551 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 552 | \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>, | |
| 553 | Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs" | |
| 554 | apply (blast dest!: ticket_authentic intro!: lemma_A) | |
| 555 | done | |
| 556 | ||
| 557 | lemma A_authenticates_B: | |
| 558 | "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs); | |
| 559 | Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | |
| 560 | Key K \<notin> analz (spies evs); | |
| 561 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 562 | \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs" | |
| 563 | apply (blast dest!: Kab_authentic intro!: lemma_B) | |
| 564 | done | |
| 565 | ||
| 566 | subsection{*Temporal guarantees, relying on a temporal check that insures that
 | |
| 567 | no oops event occurred. These are available in the sense of goal availability*} | |
| 568 | ||
| 569 | ||
| 570 | text{*Temporal treatment of confidentiality*}
 | |
| 571 | ||
| 572 | text{* Lemma: the session key sent in msg BK2 would be EXPIRED
 | |
| 573 | if the spy could see it! *} | |
| 574 | lemma lemma_conf_temporal [rule_format (no_asm)]: | |
| 575 | "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 576 | \<Longrightarrow> Says Server A | |
| 577 | (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, | |
| 578 | Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>) | |
| 579 | \<in> set evs \<longrightarrow> | |
| 580 | Key K \<in> analz (spies evs) \<longrightarrow> expiredK Tk evs" | |
| 581 | apply (erule bankerberos.induct) | |
| 13926 | 582 | apply (frule_tac [7] Says_Server_message_form) | 
| 583 | apply (frule_tac [5] Says_S_message_form [THEN disjE]) | |
| 584 | apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes) | |
| 585 | txt{*Fake*}
 | |
| 586 | apply spy_analz | |
| 18886 | 587 | txt{*BK2*}
 | 
| 13926 | 588 | apply (blast intro: parts_insertI less_SucI) | 
| 18886 | 589 | txt{*BK3*}
 | 
| 13926 | 590 | apply (case_tac "Aa \<in> bad") | 
| 18886 | 591 | prefer 2 apply (blast dest: Kab_authentic unique_session_keys) | 
| 13926 | 592 | apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI) | 
| 18886 | 593 | txt{*Oops: PROOF FAILS if unsafe intro below*}
 | 
| 13926 | 594 | apply (blast dest: unique_session_keys intro!: less_SucI) | 
| 595 | done | |
| 5053 | 596 | |
| 597 | ||
| 18886 | 598 | text{*Confidentiality for the Server: Spy does not see the keys sent in msg BK2
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 599 | as long as they have not expired.*} | 
| 18886 | 600 | lemma Confidentiality_S_temporal: | 
| 601 | "\<lbrakk> Says Server A | |
| 602 | (Crypt K' \<lbrace>Number T, Agent B, Key K, X\<rbrace>) \<in> set evs; | |
| 603 | \<not> expiredK T evs; | |
| 604 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos | |
| 605 | \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)" | |
| 13926 | 606 | apply (frule Says_Server_message_form, assumption) | 
| 18886 | 607 | apply (blast intro: lemma_conf_temporal) | 
| 13926 | 608 | done | 
| 609 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 610 | text{*Confidentiality for Alice*}
 | 
| 18886 | 611 | lemma Confidentiality_A_temporal: | 
| 612 | "\<lbrakk> Crypt (shrK A) \<lbrace>Number T, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | |
| 613 | \<not> expiredK T evs; | |
| 614 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos | |
| 615 | \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)" | |
| 616 | apply (blast dest!: Kab_authentic Confidentiality_S_temporal) | |
| 617 | done | |
| 13926 | 618 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 619 | text{*Confidentiality for Bob*}
 | 
| 18886 | 620 | lemma Confidentiality_B_temporal: | 
| 621 | "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 622 | \<in> parts (spies evs); | 
| 18886 | 623 | \<not> expiredK Tk evs; | 
| 624 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos | |
| 625 | \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)" | |
| 626 | apply (blast dest!: ticket_authentic Confidentiality_S_temporal) | |
| 627 | done | |
| 628 | ||
| 629 | text{*Temporal treatment of authentication*}
 | |
| 630 | ||
| 631 | text{*Authentication of A to B*}
 | |
| 632 | lemma B_authenticates_A_temporal: | |
| 633 | "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs); | |
| 634 | Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> | |
| 635 | \<in> parts (spies evs); | |
| 636 | \<not> expiredK Tk evs; | |
| 637 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 638 | \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>, | |
| 639 | Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs" | |
| 640 | apply (blast dest!: ticket_authentic | |
| 641 | intro!: lemma_A | |
| 642 | elim!: Confidentiality_S_temporal [THEN [2] rev_notE]) | |
| 643 | done | |
| 644 | ||
| 645 | text{*Authentication of B to A*}
 | |
| 646 | lemma A_authenticates_B_temporal: | |
| 647 | "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs); | |
| 648 | Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> | |
| 649 | \<in> parts (spies evs); | |
| 650 | \<not> expiredK Tk evs; | |
| 651 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 652 | \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs" | |
| 653 | apply (blast dest!: Kab_authentic | |
| 654 | intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE]) | |
| 655 | done | |
| 656 | ||
| 657 | subsection{*Treatment of the key distribution goal using trace inspection. All
 | |
| 658 | guarantees are in non-temporal form, hence non available, though their temporal | |
| 659 | form is trivial to derive. These guarantees also convey a stronger form of | |
| 660 | authentication - non-injective agreement on the session key*} | |
| 13926 | 661 | |
| 5053 | 662 | |
| 18886 | 663 | lemma B_Issues_A: | 
| 664 | "\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs; | |
| 665 | Key K \<notin> analz (spies evs); | |
| 666 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 667 | \<Longrightarrow> B Issues A with (Crypt K (Number Ta)) on evs" | |
| 668 | apply (simp (no_asm) add: Issues_def) | |
| 669 | apply (rule exI) | |
| 670 | apply (rule conjI, assumption) | |
| 671 | apply (simp (no_asm)) | |
| 672 | apply (erule rev_mp) | |
| 673 | apply (erule rev_mp) | |
| 674 | apply (erule bankerberos.induct, analz_mono_contra) | |
| 675 | apply (simp_all (no_asm_simp)) | |
| 676 | txt{*fake*}
 | |
| 13926 | 677 | apply blast | 
| 18886 | 678 | txt{*K4 obviously is the non-trivial case*}
 | 
| 679 | apply (simp add: takeWhile_tail) | |
| 680 | apply (blast dest: ticket_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: A_authenticates_B_temporal) | |
| 681 | done | |
| 682 | ||
| 683 | lemma A_authenticates_and_keydist_to_B: | |
| 684 | "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs); | |
| 685 | Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs); | |
| 686 | Key K \<notin> analz (spies evs); | |
| 687 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 688 | \<Longrightarrow> B Issues A with (Crypt K (Number Ta)) on evs" | |
| 689 | apply (blast dest!: A_authenticates_B B_Issues_A) | |
| 13926 | 690 | done | 
| 691 | ||
| 692 | ||
| 18886 | 693 | lemma A_Issues_B: | 
| 694 | "\<lbrakk> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> | |
| 695 | \<in> set evs; | |
| 696 | Key K \<notin> analz (spies evs); | |
| 697 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 698 | \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) on evs" | |
| 699 | apply (simp (no_asm) add: Issues_def) | |
| 700 | apply (rule exI) | |
| 701 | apply (rule conjI, assumption) | |
| 702 | apply (simp (no_asm)) | |
| 703 | apply (erule rev_mp) | |
| 704 | apply (erule rev_mp) | |
| 705 | apply (erule bankerberos.induct, analz_mono_contra) | |
| 706 | apply (simp_all (no_asm_simp)) | |
| 707 | txt{*fake*}
 | |
| 13926 | 708 | apply blast | 
| 18886 | 709 | txt{*K3 is the non trivial case*}
 | 
| 710 | apply (simp add: takeWhile_tail) | |
| 711 | apply auto (*Technically unnecessary, merely clarifies the subgoal as it is presemted in the book*) | |
| 712 | apply (blast dest: Kab_authentic Says_Server_message_form parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] | |
| 713 | intro!: B_authenticates_A) | |
| 13926 | 714 | done | 
| 715 | ||
| 18886 | 716 | |
| 717 | lemma B_authenticates_and_keydist_to_A: | |
| 718 | "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs); | |
| 719 | Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs); | |
| 720 | Key K \<notin> analz (spies evs); | |
| 721 | A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk> | |
| 722 | \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) on evs" | |
| 723 | apply (blast dest: B_authenticates_A A_Issues_B) | |
| 724 | done | |
| 725 | ||
| 726 | ||
| 727 | ||
| 5053 | 728 | |
| 729 | end |