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