| 18886 |      1 | (*  ID:         $Id$
 | 
|  |      2 |     Author:     Giampaolo Bella, Catania University
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | header{*The Kerberos Protocol, Version V*}
 | 
|  |      6 | 
 | 
|  |      7 | theory KerberosV imports Public begin
 | 
|  |      8 | 
 | 
|  |      9 | text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
 | 
|  |     10 | 
 | 
|  |     11 | syntax
 | 
|  |     12 |   Kas :: agent
 | 
|  |     13 |   Tgs :: agent  --{*the two servers are translations...*}
 | 
|  |     14 | 
 | 
|  |     15 | 
 | 
|  |     16 | translations
 | 
|  |     17 |   "Kas"       == "Server "
 | 
|  |     18 |   "Tgs"       == "Friend 0"
 | 
|  |     19 | 
 | 
|  |     20 | 
 | 
|  |     21 | axioms
 | 
|  |     22 |   Tgs_not_bad [iff]: "Tgs \<notin> bad"
 | 
|  |     23 |    --{*Tgs is secure --- we already know that Kas is secure*}
 | 
|  |     24 | 
 | 
|  |     25 | (*The current time is just the length of the trace!*)
 | 
|  |     26 | syntax
 | 
|  |     27 |     CT :: "event list=>nat"
 | 
|  |     28 | 
 | 
|  |     29 |     expiredAK :: "[nat, event list] => bool"
 | 
|  |     30 | 
 | 
|  |     31 |     expiredSK :: "[nat, event list] => bool"
 | 
|  |     32 | 
 | 
|  |     33 |     expiredA :: "[nat, event list] => bool"
 | 
|  |     34 | 
 | 
|  |     35 |     valid :: "[nat, nat] => bool"  ("valid _ wrt _")
 | 
|  |     36 | 
 | 
|  |     37 | 
 | 
|  |     38 | constdefs
 | 
|  |     39 |  (* authKeys are those contained in an authTicket *)
 | 
|  |     40 |     authKeys :: "event list => key set"
 | 
|  |     41 |     "authKeys evs == {authK. \<exists>A Peer Ta. 
 | 
|  |     42 |         Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Peer, Ta\<rbrace>,
 | 
|  |     43 |                      Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key authK, Ta\<rbrace>
 | 
|  |     44 |                    \<rbrace> \<in> set evs}"
 | 
|  |     45 | 
 | 
|  |     46 |  (* A is the true creator of X if she has sent X and X never appeared on
 | 
|  |     47 |     the trace before this event. Recall that traces grow from head. *)
 | 
|  |     48 |   Issues :: "[agent, agent, msg, event list] => bool"
 | 
|  |     49 |              ("_ Issues _ with _ on _")
 | 
|  |     50 |    "A Issues B with X on evs ==
 | 
|  |     51 |       \<exists>Y. Says A B Y \<in> set evs \<and> X \<in> parts {Y} \<and>
 | 
|  |     52 |       X \<notin> parts (spies (takeWhile (% z. z  \<noteq> Says A B Y) (rev evs)))"
 | 
|  |     53 | 
 | 
|  |     54 | 
 | 
|  |     55 | consts
 | 
|  |     56 |     (*Duration of the authentication key*)
 | 
|  |     57 |     authKlife   :: nat
 | 
|  |     58 | 
 | 
|  |     59 |     (*Duration of the service key*)
 | 
|  |     60 |     servKlife   :: nat
 | 
|  |     61 | 
 | 
|  |     62 |     (*Duration of an authenticator*)
 | 
|  |     63 |     authlife   :: nat
 | 
|  |     64 | 
 | 
|  |     65 |     (*Upper bound on the time of reaction of a server*)
 | 
|  |     66 |     replylife   :: nat
 | 
|  |     67 | 
 | 
|  |     68 | specification (authKlife)
 | 
|  |     69 |   authKlife_LB [iff]: "2 \<le> authKlife"
 | 
|  |     70 |     by blast
 | 
|  |     71 | 
 | 
|  |     72 | specification (servKlife)
 | 
|  |     73 |   servKlife_LB [iff]: "2 + authKlife \<le> servKlife"
 | 
|  |     74 |     by blast
 | 
|  |     75 | 
 | 
|  |     76 | specification (authlife)
 | 
|  |     77 |   authlife_LB [iff]: "Suc 0 \<le> authlife"
 | 
|  |     78 |     by blast
 | 
|  |     79 | 
 | 
|  |     80 | specification (replylife)
 | 
|  |     81 |   replylife_LB [iff]: "Suc 0 \<le> replylife"
 | 
|  |     82 |     by blast
 | 
|  |     83 | 
 | 
|  |     84 | translations
 | 
|  |     85 |    "CT" == "length "
 | 
|  |     86 | 
 | 
|  |     87 |    "expiredAK T evs" == "authKlife + T < CT evs"
 | 
|  |     88 | 
 | 
|  |     89 |    "expiredSK T evs" == "servKlife + T < CT evs"
 | 
|  |     90 | 
 | 
|  |     91 |    "expiredA T evs" == "authlife + T < CT evs"
 | 
|  |     92 | 
 | 
|  |     93 |    "valid T1 wrt T2" == "T1 <= replylife + T2"
 | 
|  |     94 | 
 | 
|  |     95 | (*---------------------------------------------------------------------*)
 | 
|  |     96 | 
 | 
|  |     97 | 
 | 
|  |     98 | (* Predicate formalising the association between authKeys and servKeys *)
 | 
|  |     99 | constdefs
 | 
|  |    100 |   AKcryptSK :: "[key, key, event list] => bool"
 | 
|  |    101 |   "AKcryptSK authK servK evs ==
 | 
|  |    102 |      \<exists>A B tt.
 | 
|  |    103 |        Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>,
 | 
|  |    104 |                     Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, tt\<rbrace> \<rbrace>
 | 
|  |    105 |          \<in> set evs"
 | 
|  |    106 | 
 | 
|  |    107 | consts
 | 
|  |    108 | 
 | 
|  |    109 | kerbV   :: "event list set"
 | 
|  |    110 | inductive "kerbV"
 | 
|  |    111 |   intros
 | 
|  |    112 | 
 | 
|  |    113 |    Nil:  "[] \<in> kerbV"
 | 
|  |    114 | 
 | 
|  |    115 |    Fake: "\<lbrakk> evsf \<in> kerbV;  X \<in> synth (analz (spies evsf)) \<rbrakk>
 | 
|  |    116 |           \<Longrightarrow> Says Spy B X  # evsf \<in> kerbV"
 | 
|  |    117 | 
 | 
|  |    118 | 
 | 
|  |    119 | (*Authentication phase*)
 | 
|  |    120 |    KV1:   "\<lbrakk> evs1 \<in> kerbV \<rbrakk>
 | 
|  |    121 |           \<Longrightarrow> Says A Kas \<lbrace>Agent A, Agent Tgs, Number (CT evs1)\<rbrace> # evs1
 | 
|  |    122 |           \<in> kerbV"
 | 
|  |    123 |    (*Unlike version IV, authTicket is not re-encrypted*)
 | 
|  |    124 |    KV2:  "\<lbrakk> evs2 \<in> kerbV; Key authK \<notin> used evs2; authK \<in> symKeys;
 | 
|  |    125 |             Says A' Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs2 \<rbrakk>
 | 
|  |    126 |           \<Longrightarrow> Says Kas A \<lbrace>
 | 
|  |    127 |           Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number (CT evs2)\<rbrace>,
 | 
|  |    128 |         Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number (CT evs2)\<rbrace> 
 | 
|  |    129 |                          \<rbrace> # evs2 \<in> kerbV"
 | 
|  |    130 | 
 | 
|  |    131 | 
 | 
|  |    132 | (* Authorisation phase *)
 | 
|  |    133 |    KV3:  "\<lbrakk> evs3 \<in> kerbV; A \<noteq> Kas; A \<noteq> Tgs;
 | 
|  |    134 |             Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs3;
 | 
|  |    135 |             Says Kas' A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    136 |                           authTicket\<rbrace> \<in> set evs3;
 | 
|  |    137 |             valid Ta wrt T1
 | 
|  |    138 |          \<rbrakk>
 | 
|  |    139 |           \<Longrightarrow> Says A Tgs \<lbrace>authTicket,
 | 
|  |    140 |                            (Crypt authK \<lbrace>Agent A, Number (CT evs3)\<rbrace>),
 | 
|  |    141 |                            Agent B\<rbrace> # evs3 \<in> kerbV"
 | 
|  |    142 |    (*Unlike version IV, servTicket is not re-encrypted*)
 | 
|  |    143 |    KV4:  "\<lbrakk> evs4 \<in> kerbV; Key servK \<notin> used evs4; servK \<in> symKeys;
 | 
|  |    144 |             B \<noteq> Tgs;  authK \<in> symKeys;
 | 
|  |    145 |             Says A' Tgs \<lbrace>
 | 
|  |    146 |              (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
 | 
|  |    147 | 				 Number Ta\<rbrace>),
 | 
|  |    148 |              (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
 | 
|  |    149 | 	        \<in> set evs4;
 | 
|  |    150 |             \<not> expiredAK Ta evs4;
 | 
|  |    151 |             \<not> expiredA T2 evs4;
 | 
|  |    152 |             servKlife + (CT evs4) <= authKlife + Ta
 | 
|  |    153 |          \<rbrakk>
 | 
|  |    154 |           \<Longrightarrow> Says Tgs A \<lbrace>
 | 
|  |    155 |              Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4)\<rbrace>,
 | 
|  |    156 |    Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number (CT evs4)\<rbrace> 
 | 
|  |    157 |                           \<rbrace> # evs4 \<in> kerbV"
 | 
|  |    158 | 
 | 
|  |    159 | 
 | 
|  |    160 | (*Service phase*)
 | 
|  |    161 |    KV5:  "\<lbrakk> evs5 \<in> kerbV; authK \<in> symKeys; servK \<in> symKeys;
 | 
|  |    162 |             A \<noteq> Kas; A \<noteq> Tgs;
 | 
|  |    163 |             Says A Tgs
 | 
|  |    164 |                 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
 | 
|  |    165 | 		  Agent B\<rbrace>
 | 
|  |    166 |               \<in> set evs5;
 | 
|  |    167 |             Says Tgs' A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
 | 
|  |    168 |                           servTicket\<rbrace>
 | 
|  |    169 |                 \<in> set evs5;
 | 
|  |    170 |             valid Ts wrt T2 \<rbrakk>
 | 
|  |    171 |           \<Longrightarrow> Says A B \<lbrace>servTicket,
 | 
|  |    172 | 			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
 | 
|  |    173 |                # evs5 \<in> kerbV"
 | 
|  |    174 | 
 | 
|  |    175 |     KV6:  "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
 | 
|  |    176 |             Says A' B \<lbrace>
 | 
|  |    177 |               (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>),
 | 
|  |    178 |               (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>)\<rbrace>
 | 
|  |    179 |             \<in> set evs6;
 | 
|  |    180 |             \<not> expiredSK Ts evs6;
 | 
|  |    181 |             \<not> expiredA T3 evs6
 | 
|  |    182 |          \<rbrakk>
 | 
|  |    183 |           \<Longrightarrow> Says B A (Crypt servK (Number Ta2))
 | 
|  |    184 |                # evs6 \<in> kerbV"
 | 
|  |    185 | 
 | 
|  |    186 | 
 | 
|  |    187 | 
 | 
|  |    188 | (* Leaking an authK... *)
 | 
|  |    189 |    Oops1:"\<lbrakk> evsO1 \<in> kerbV;  A \<noteq> Spy;
 | 
|  |    190 |              Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    191 |                           authTicket\<rbrace>  \<in> set evsO1;
 | 
|  |    192 |               expiredAK Ta evsO1 \<rbrakk>
 | 
|  |    193 |           \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent Tgs, Number Ta, Key authK\<rbrace>
 | 
|  |    194 |                # evsO1 \<in> kerbV"
 | 
|  |    195 | 
 | 
|  |    196 | (*Leaking a servK... *)
 | 
|  |    197 |    Oops2: "\<lbrakk> evsO2 \<in> kerbV;  A \<noteq> Spy;
 | 
|  |    198 |               Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
 | 
|  |    199 |                            servTicket\<rbrace>  \<in> set evsO2;
 | 
|  |    200 |               expiredSK Ts evsO2 \<rbrakk>
 | 
|  |    201 |           \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Number Ts, Key servK\<rbrace>
 | 
|  |    202 |                # evsO2 \<in> kerbV"
 | 
|  |    203 | 
 | 
|  |    204 | 
 | 
|  |    205 | 
 | 
|  |    206 | declare Says_imp_knows_Spy [THEN parts.Inj, dest]
 | 
|  |    207 | declare parts.Body [dest]
 | 
|  |    208 | declare analz_into_parts [dest]
 | 
|  |    209 | declare Fake_parts_insert_in_Un [dest]
 | 
|  |    210 | 
 | 
|  |    211 | 
 | 
|  |    212 | 
 | 
|  |    213 | subsection{*Lemmas about lists, for reasoning about  Issues*}
 | 
|  |    214 | 
 | 
|  |    215 | lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 | 
|  |    216 | apply (induct_tac "evs")
 | 
|  |    217 | apply (induct_tac [2] "a", auto)
 | 
|  |    218 | done
 | 
|  |    219 | 
 | 
|  |    220 | lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
 | 
|  |    221 | apply (induct_tac "evs")
 | 
|  |    222 | apply (induct_tac [2] "a", auto)
 | 
|  |    223 | done
 | 
|  |    224 | 
 | 
|  |    225 | lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
 | 
|  |    226 |           (if A:bad then insert X (spies evs) else spies evs)"
 | 
|  |    227 | apply (induct_tac "evs")
 | 
|  |    228 | apply (induct_tac [2] "a", auto)
 | 
|  |    229 | done
 | 
|  |    230 | 
 | 
|  |    231 | lemma spies_evs_rev: "spies evs = spies (rev evs)"
 | 
|  |    232 | apply (induct_tac "evs")
 | 
|  |    233 | apply (induct_tac [2] "a")
 | 
|  |    234 | apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
 | 
|  |    235 | done
 | 
|  |    236 | 
 | 
|  |    237 | lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono]
 | 
|  |    238 | 
 | 
|  |    239 | lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
 | 
|  |    240 | apply (induct_tac "evs")
 | 
|  |    241 | apply (induct_tac [2] "a", auto)
 | 
|  |    242 | txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 | 
|  |    243 | done
 | 
|  |    244 | 
 | 
|  |    245 | lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
 | 
|  |    246 | 
 | 
|  |    247 | 
 | 
|  |    248 | subsection{*Lemmas about @{term authKeys}*}
 | 
|  |    249 | 
 | 
|  |    250 | lemma authKeys_empty: "authKeys [] = {}"
 | 
|  |    251 | apply (unfold authKeys_def)
 | 
|  |    252 | apply (simp (no_asm))
 | 
|  |    253 | done
 | 
|  |    254 | 
 | 
|  |    255 | lemma authKeys_not_insert:
 | 
|  |    256 |  "(\<forall>A Ta akey Peer.
 | 
|  |    257 |    ev \<noteq> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>akey, Agent Peer, Ta\<rbrace>,
 | 
|  |    258 |                      Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, akey, Ta\<rbrace> \<rbrace>)
 | 
|  |    259 |        \<Longrightarrow> authKeys (ev # evs) = authKeys evs"
 | 
|  |    260 | apply (unfold authKeys_def, auto)
 | 
|  |    261 | done
 | 
|  |    262 | 
 | 
|  |    263 | lemma authKeys_insert:
 | 
|  |    264 |   "authKeys
 | 
|  |    265 |      (Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key K, Agent Peer, Number Ta\<rbrace>,
 | 
|  |    266 |          Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K, Number Ta\<rbrace> \<rbrace> # evs)
 | 
|  |    267 |        = insert K (authKeys evs)"
 | 
|  |    268 | apply (unfold authKeys_def, auto)
 | 
|  |    269 | done
 | 
|  |    270 | 
 | 
|  |    271 | lemma authKeys_simp:
 | 
|  |    272 |    "K \<in> authKeys
 | 
|  |    273 |     (Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key K', Agent Peer, Number Ta\<rbrace>,
 | 
|  |    274 |         Crypt (shrK Peer) \<lbrace>Agent A, Agent Peer, Key K', Number Ta\<rbrace> \<rbrace> # evs)
 | 
|  |    275 |         \<Longrightarrow> K = K' | K \<in> authKeys evs"
 | 
|  |    276 | apply (unfold authKeys_def, auto)
 | 
|  |    277 | done
 | 
|  |    278 | 
 | 
|  |    279 | lemma authKeysI:
 | 
|  |    280 |    "Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key K, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    281 |          Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key K, Number Ta\<rbrace> \<rbrace> \<in> set evs
 | 
|  |    282 |         \<Longrightarrow> K \<in> authKeys evs"
 | 
|  |    283 | apply (unfold authKeys_def, auto)
 | 
|  |    284 | done
 | 
|  |    285 | 
 | 
|  |    286 | lemma authKeys_used: "K \<in> authKeys evs \<Longrightarrow> Key K \<in> used evs"
 | 
|  |    287 | apply (simp add: authKeys_def, blast)
 | 
|  |    288 | done
 | 
|  |    289 | 
 | 
|  |    290 | 
 | 
|  |    291 | subsection{*Forwarding Lemmas*}
 | 
|  |    292 | 
 | 
|  |    293 | lemma Says_ticket_parts:
 | 
|  |    294 |      "Says S A \<lbrace>Crypt K \<lbrace>SesKey, B, TimeStamp\<rbrace>, Ticket\<rbrace>
 | 
|  |    295 |                \<in> set evs \<Longrightarrow> Ticket \<in> parts (spies evs)"
 | 
|  |    296 | apply blast
 | 
|  |    297 | done
 | 
|  |    298 | 
 | 
|  |    299 | lemma Says_ticket_analz:
 | 
|  |    300 |      "Says S A \<lbrace>Crypt K \<lbrace>SesKey, B, TimeStamp\<rbrace>, Ticket\<rbrace>
 | 
|  |    301 |                \<in> set evs \<Longrightarrow> Ticket \<in> analz (spies evs)"
 | 
|  |    302 | apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd])
 | 
|  |    303 | done
 | 
|  |    304 | 
 | 
|  |    305 | lemma Oops_range_spies1:
 | 
|  |    306 |      "\<lbrakk> Says Kas A \<lbrace>Crypt KeyA \<lbrace>Key authK, Peer, Ta\<rbrace>, authTicket\<rbrace>
 | 
|  |    307 |            \<in> set evs ;
 | 
|  |    308 |          evs \<in> kerbV \<rbrakk> \<Longrightarrow> authK \<notin> range shrK & authK \<in> symKeys"
 | 
|  |    309 | apply (erule rev_mp)
 | 
|  |    310 | apply (erule kerbV.induct, auto)
 | 
|  |    311 | done
 | 
|  |    312 | 
 | 
|  |    313 | lemma Oops_range_spies2:
 | 
|  |    314 |      "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |    315 |            \<in> set evs ;
 | 
|  |    316 |          evs \<in> kerbV \<rbrakk> \<Longrightarrow> servK \<notin> range shrK \<and> servK \<in> symKeys"
 | 
|  |    317 | apply (erule rev_mp)
 | 
|  |    318 | apply (erule kerbV.induct, auto)
 | 
|  |    319 | done
 | 
|  |    320 | 
 | 
|  |    321 | 
 | 
|  |    322 | (*Spy never sees another agent's shared key! (unless it's lost at start)*)
 | 
|  |    323 | lemma Spy_see_shrK [simp]:
 | 
|  |    324 |      "evs \<in> kerbV \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
 | 
|  |    325 | apply (erule kerbV.induct)
 | 
|  |    326 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    327 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    328 | apply (blast+)
 | 
|  |    329 | done
 | 
|  |    330 | 
 | 
|  |    331 | lemma Spy_analz_shrK [simp]:
 | 
|  |    332 |      "evs \<in> kerbV \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
 | 
|  |    333 | by auto
 | 
|  |    334 | 
 | 
|  |    335 | lemma Spy_see_shrK_D [dest!]:
 | 
|  |    336 |      "\<lbrakk> Key (shrK A) \<in> parts (spies evs);  evs \<in> kerbV \<rbrakk> \<Longrightarrow> A:bad"
 | 
|  |    337 | by (blast dest: Spy_see_shrK)
 | 
|  |    338 | lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!]
 | 
|  |    339 | 
 | 
|  |    340 | text{*Nobody can have used non-existent keys!*}
 | 
|  |    341 | lemma new_keys_not_used [simp]:
 | 
|  |    342 |     "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerbV\<rbrakk>
 | 
|  |    343 |      \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
 | 
|  |    344 | apply (erule rev_mp)
 | 
|  |    345 | apply (erule kerbV.induct)
 | 
|  |    346 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    347 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    348 | txt{*Fake*}
 | 
|  |    349 | apply (force dest!: keysFor_parts_insert)
 | 
|  |    350 | txt{*Others*}
 | 
|  |    351 | apply (force dest!: analz_shrK_Decrypt)+
 | 
|  |    352 | done
 | 
|  |    353 | 
 | 
|  |    354 | (*Earlier, all protocol proofs declared this theorem.
 | 
|  |    355 |   But few of them actually need it! (Another is Yahalom) *)
 | 
|  |    356 | lemma new_keys_not_analzd:
 | 
|  |    357 |  "\<lbrakk>evs \<in> kerbV; K \<in> symKeys; Key K \<notin> used evs\<rbrakk>
 | 
|  |    358 |   \<Longrightarrow> K \<notin> keysFor (analz (spies evs))"
 | 
|  |    359 | by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD])
 | 
|  |    360 | 
 | 
|  |    361 | 
 | 
|  |    362 | 
 | 
|  |    363 | subsection{*Regularity Lemmas*}
 | 
|  |    364 | text{*These concern the form of items passed in messages*}
 | 
|  |    365 | 
 | 
|  |    366 | text{*Describes the form of all components sent by Kas*}
 | 
|  |    367 | lemma Says_Kas_message_form:
 | 
|  |    368 |      "\<lbrakk> Says Kas A \<lbrace>Crypt K \<lbrace>Key authK, Agent Peer, Ta\<rbrace>, authTicket\<rbrace>
 | 
|  |    369 |            \<in> set evs;
 | 
|  |    370 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    371 |       \<Longrightarrow> authK \<notin> range shrK \<and> authK \<in> authKeys evs \<and> authK \<in> symKeys \<and> 
 | 
|  |    372 |   authTicket = (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>) \<and>
 | 
|  |    373 |              K = shrK A  \<and> Peer = Tgs"
 | 
|  |    374 | apply (erule rev_mp)
 | 
|  |    375 | apply (erule kerbV.induct)
 | 
|  |    376 | apply (simp_all (no_asm) add: authKeys_def authKeys_insert)
 | 
|  |    377 | apply blast+
 | 
|  |    378 | done
 | 
|  |    379 | 
 | 
|  |    380 | 
 | 
|  |    381 | 
 | 
|  |    382 | (*This lemma is essential for proving Says_Tgs_message_form:
 | 
|  |    383 | 
 | 
|  |    384 |   the session key authK
 | 
|  |    385 |   supplied by Kas in the authentication ticket
 | 
|  |    386 |   cannot be a long-term key!
 | 
|  |    387 | 
 | 
|  |    388 |   Generalised to any session keys (both authK and servK).
 | 
|  |    389 | *)
 | 
|  |    390 | lemma SesKey_is_session_key:
 | 
|  |    391 |      "\<lbrakk> Crypt (shrK Tgs_B) \<lbrace>Agent A, Agent Tgs_B, Key SesKey, Number T\<rbrace>
 | 
|  |    392 |             \<in> parts (spies evs); Tgs_B \<notin> bad;
 | 
|  |    393 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    394 |       \<Longrightarrow> SesKey \<notin> range shrK"
 | 
|  |    395 | apply (erule rev_mp)
 | 
|  |    396 | apply (erule kerbV.induct)
 | 
|  |    397 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    398 | apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
 | 
|  |    399 | done
 | 
|  |    400 | 
 | 
|  |    401 | lemma authTicket_authentic:
 | 
|  |    402 |      "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>
 | 
|  |    403 |            \<in> parts (spies evs);
 | 
|  |    404 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    405 |       \<Longrightarrow> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>,
 | 
|  |    406 |                  Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Ta\<rbrace>\<rbrace>
 | 
|  |    407 |             \<in> set evs"
 | 
|  |    408 | apply (erule rev_mp)
 | 
|  |    409 | apply (erule kerbV.induct)
 | 
|  |    410 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    411 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    412 | txt{*Fake, K4*}
 | 
|  |    413 | apply (blast+)
 | 
|  |    414 | done
 | 
|  |    415 | 
 | 
|  |    416 | lemma authTicket_crypt_authK:
 | 
|  |    417 |      "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>
 | 
|  |    418 |            \<in> parts (spies evs);
 | 
|  |    419 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    420 |       \<Longrightarrow> authK \<in> authKeys evs"
 | 
|  |    421 | apply (frule authTicket_authentic, assumption)
 | 
|  |    422 | apply (simp (no_asm) add: authKeys_def)
 | 
|  |    423 | apply blast
 | 
|  |    424 | done
 | 
|  |    425 | 
 | 
|  |    426 | text{*Describes the form of servK, servTicket and authK sent by Tgs*}
 | 
|  |    427 | lemma Says_Tgs_message_form:
 | 
|  |    428 |      "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |    429 |            \<in> set evs;
 | 
|  |    430 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    431 |    \<Longrightarrow> B \<noteq> Tgs \<and> 
 | 
|  |    432 |        servK \<notin> range shrK \<and> servK \<notin> authKeys evs \<and> servK \<in> symKeys \<and>
 | 
|  |    433 |        servTicket = (Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>) \<and>
 | 
|  |    434 |        authK \<notin> range shrK \<and> authK \<in> authKeys evs \<and> authK \<in> symKeys"
 | 
|  |    435 | apply (erule rev_mp)
 | 
|  |    436 | apply (erule kerbV.induct)
 | 
|  |    437 | apply (simp_all add: authKeys_insert authKeys_not_insert authKeys_empty authKeys_simp, blast, auto)
 | 
|  |    438 | txt{*Three subcases of Message 4*}
 | 
|  |    439 | apply (blast dest!: authKeys_used Says_Kas_message_form)
 | 
|  |    440 | apply (blast dest!: SesKey_is_session_key)
 | 
|  |    441 | apply (blast dest: authTicket_crypt_authK)
 | 
|  |    442 | done
 | 
|  |    443 | 
 | 
|  |    444 | 
 | 
|  |    445 | 
 | 
|  |    446 | (*
 | 
|  |    447 | lemma authTicket_form:
 | 
|  |    448 | lemma servTicket_form:
 | 
|  |    449 | lemma Says_kas_message_form:
 | 
|  |    450 | lemma Says_tgs_message_form:
 | 
|  |    451 | 
 | 
|  |    452 | cannot be proved for version V, but a new proof strategy can be used in their
 | 
|  |    453 | place. The new strategy merely says that both the authTicket and the servTicket
 | 
|  |    454 | are in parts and in analz as soon as they appear, using lemmas Says_ticket_parts and Says_ticket_analz. 
 | 
|  |    455 | The new strategy always lets the simplifier solve cases K3 and K5, saving on
 | 
|  |    456 | long dedicated analyses, which seemed unavoidable. For this reason, lemma 
 | 
|  |    457 | servK_notin_authKeysD is no longer needed.
 | 
|  |    458 | *)
 | 
|  |    459 | 
 | 
|  |    460 | subsection{*Authenticity theorems: confirm origin of sensitive messages*}
 | 
|  |    461 | 
 | 
|  |    462 | lemma authK_authentic:
 | 
|  |    463 |      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>
 | 
|  |    464 |            \<in> parts (spies evs);
 | 
|  |    465 |          A \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |    466 |       \<Longrightarrow> \<exists> AT. Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>, AT\<rbrace>
 | 
|  |    467 |             \<in> set evs"
 | 
|  |    468 | apply (erule rev_mp)
 | 
|  |    469 | apply (erule kerbV.induct)
 | 
|  |    470 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    471 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    472 | apply blast+
 | 
|  |    473 | done
 | 
|  |    474 | 
 | 
|  |    475 | text{*If a certain encrypted message appears then it originated with Tgs*}
 | 
|  |    476 | lemma servK_authentic:
 | 
|  |    477 |      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>
 | 
|  |    478 |            \<in> parts (spies evs);
 | 
|  |    479 |          Key authK \<notin> analz (spies evs);
 | 
|  |    480 |          authK \<notin> range shrK;
 | 
|  |    481 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    482 |  \<Longrightarrow> \<exists>A ST. Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, ST\<rbrace>
 | 
|  |    483 |        \<in> set evs"
 | 
|  |    484 | apply (erule rev_mp)
 | 
|  |    485 | apply (erule rev_mp)
 | 
|  |    486 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    487 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    488 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    489 | apply blast+
 | 
|  |    490 | done
 | 
|  |    491 | 
 | 
|  |    492 | lemma servK_authentic_bis:
 | 
|  |    493 |      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>
 | 
|  |    494 |            \<in> parts (spies evs);
 | 
|  |    495 |          Key authK \<notin> analz (spies evs);
 | 
|  |    496 |          B \<noteq> Tgs;
 | 
|  |    497 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    498 |  \<Longrightarrow> \<exists>A ST. Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, ST\<rbrace>
 | 
|  |    499 |        \<in> set evs"
 | 
|  |    500 | apply (erule rev_mp)
 | 
|  |    501 | apply (erule rev_mp)
 | 
|  |    502 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    503 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    504 | apply (frule_tac [5] Says_ticket_parts, simp_all, blast+)
 | 
|  |    505 | done
 | 
|  |    506 | 
 | 
|  |    507 | text{*Authenticity of servK for B*}
 | 
|  |    508 | lemma servTicket_authentic_Tgs:
 | 
|  |    509 |      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>
 | 
|  |    510 |            \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
 | 
|  |    511 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    512 |  \<Longrightarrow> \<exists>authK.
 | 
|  |    513 |        Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>,
 | 
|  |    514 |                     Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>\<rbrace>
 | 
|  |    515 |        \<in> set evs"
 | 
|  |    516 | apply (erule rev_mp)
 | 
|  |    517 | apply (erule kerbV.induct)
 | 
|  |    518 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    519 | apply (frule_tac [5] Says_ticket_parts, simp_all, blast+)
 | 
|  |    520 | done
 | 
|  |    521 | 
 | 
|  |    522 | text{*Anticipated here from next subsection*}
 | 
|  |    523 | lemma K4_imp_K2:
 | 
|  |    524 | "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |    525 |       \<in> set evs;  evs \<in> kerbV\<rbrakk>
 | 
|  |    526 |    \<Longrightarrow> \<exists>Ta. Says Kas A
 | 
|  |    527 |         \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    528 |            Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
 | 
|  |    529 |         \<in> set evs"
 | 
|  |    530 | apply (erule rev_mp)
 | 
|  |    531 | apply (erule kerbV.induct)
 | 
|  |    532 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    533 | apply (frule_tac [5] Says_ticket_parts, simp_all, auto)
 | 
|  |    534 | apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
 | 
|  |    535 | done
 | 
|  |    536 | 
 | 
|  |    537 | text{*Anticipated here from next subsection*}
 | 
|  |    538 | lemma u_K4_imp_K2:
 | 
|  |    539 | "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>  \<in> set evs; evs \<in> kerbV\<rbrakk>
 | 
|  |    540 |    \<Longrightarrow> \<exists>Ta. Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    541 |              Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
 | 
|  |    542 |              \<in> set evs
 | 
|  |    543 |           \<and> servKlife + Ts <= authKlife + Ta"
 | 
|  |    544 | apply (erule rev_mp)
 | 
|  |    545 | apply (erule kerbV.induct)
 | 
|  |    546 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    547 | apply (frule_tac [5] Says_ticket_parts, simp_all, auto)
 | 
|  |    548 | apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authTicket_authentic])
 | 
|  |    549 | done
 | 
|  |    550 | 
 | 
|  |    551 | lemma servTicket_authentic_Kas:
 | 
|  |    552 |      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |    553 |            \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
 | 
|  |    554 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    555 |   \<Longrightarrow> \<exists>authK Ta.
 | 
|  |    556 |        Says Kas A
 | 
|  |    557 |          \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    558 |            Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
 | 
|  |    559 |         \<in> set evs"
 | 
|  |    560 | apply (blast dest!: servTicket_authentic_Tgs K4_imp_K2)
 | 
|  |    561 | done
 | 
|  |    562 | 
 | 
|  |    563 | lemma u_servTicket_authentic_Kas:
 | 
|  |    564 |      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |    565 |            \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
 | 
|  |    566 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    567 |   \<Longrightarrow> \<exists>authK Ta.
 | 
|  |    568 |        Says Kas A
 | 
|  |    569 |          \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    570 |            Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>
 | 
|  |    571 |         \<in> set evs \<and> 
 | 
|  |    572 |       servKlife + Ts <= authKlife + Ta"
 | 
|  |    573 | apply (blast dest!: servTicket_authentic_Tgs u_K4_imp_K2)
 | 
|  |    574 | done
 | 
|  |    575 | 
 | 
|  |    576 | lemma servTicket_authentic:
 | 
|  |    577 |      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |    578 |            \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
 | 
|  |    579 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    580 |  \<Longrightarrow> \<exists>Ta authK.
 | 
|  |    581 |      Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    582 |                   Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace> \<rbrace>  \<in> set evs
 | 
|  |    583 |      \<and> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
 | 
|  |    584 |                   Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
 | 
|  |    585 |        \<in> set evs"
 | 
|  |    586 | apply (blast dest: servTicket_authentic_Tgs K4_imp_K2)
 | 
|  |    587 | done
 | 
|  |    588 | 
 | 
|  |    589 | lemma u_servTicket_authentic:
 | 
|  |    590 |      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |    591 |            \<in> parts (spies evs);  B \<noteq> Tgs;  B \<notin> bad;
 | 
|  |    592 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    593 |  \<Longrightarrow> \<exists>Ta authK.
 | 
|  |    594 |      Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |    595 |                   Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace> \<in> set evs
 | 
|  |    596 |      \<and> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
 | 
|  |    597 |                  Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
 | 
|  |    598 |        \<in> set evs
 | 
|  |    599 |      \<and> servKlife + Ts <= authKlife + Ta"
 | 
|  |    600 | apply (blast dest: servTicket_authentic_Tgs u_K4_imp_K2)
 | 
|  |    601 | done
 | 
|  |    602 | 
 | 
|  |    603 | lemma u_NotexpiredSK_NotexpiredAK:
 | 
|  |    604 |      "\<lbrakk> \<not> expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \<rbrakk>
 | 
|  |    605 |       \<Longrightarrow> \<not> expiredAK Ta evs"
 | 
|  |    606 | apply (blast dest: leI le_trans dest: leD)
 | 
|  |    607 | done
 | 
|  |    608 | 
 | 
|  |    609 | 
 | 
|  |    610 | subsection{* Reliability: friendly agents send somthing if something else happened*}
 | 
|  |    611 | 
 | 
|  |    612 | lemma K3_imp_K2:
 | 
|  |    613 |      "\<lbrakk> Says A Tgs
 | 
|  |    614 |              \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace>
 | 
|  |    615 |            \<in> set evs;
 | 
|  |    616 |          A \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |    617 |       \<Longrightarrow> \<exists>Ta AT. Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>, 
 | 
|  |    618 |                                AT\<rbrace> \<in> set evs"
 | 
|  |    619 | apply (erule rev_mp)
 | 
|  |    620 | apply (erule kerbV.induct)
 | 
|  |    621 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    622 | apply (frule_tac [5] Says_ticket_parts, simp_all, blast, blast)
 | 
|  |    623 | apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN authK_authentic])
 | 
|  |    624 | done
 | 
|  |    625 | 
 | 
|  |    626 | text{*Anticipated here from next subsection. An authK is encrypted by one and only one Shared key. A servK is encrypted by one and only one authK.*}
 | 
|  |    627 | lemma Key_unique_SesKey:
 | 
|  |    628 |      "\<lbrakk> Crypt K  \<lbrace>Key SesKey,  Agent B, T\<rbrace>
 | 
|  |    629 |            \<in> parts (spies evs);
 | 
|  |    630 |          Crypt K' \<lbrace>Key SesKey,  Agent B', T'\<rbrace>
 | 
|  |    631 |            \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
 | 
|  |    632 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    633 |       \<Longrightarrow> K=K' \<and> B=B' \<and> T=T'"
 | 
|  |    634 | apply (erule rev_mp)
 | 
|  |    635 | apply (erule rev_mp)
 | 
|  |    636 | apply (erule rev_mp)
 | 
|  |    637 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    638 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    639 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    640 | txt{*Fake, K2, K4*}
 | 
|  |    641 | apply (blast+)
 | 
|  |    642 | done
 | 
|  |    643 | 
 | 
|  |    644 | text{*This inevitably has an existential form in version V*}
 | 
|  |    645 | lemma Says_K5:
 | 
|  |    646 |      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
 | 
|  |    647 |          Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
 | 
|  |    648 |                                      servTicket\<rbrace> \<in> set evs;
 | 
|  |    649 |          Key servK \<notin> analz (spies evs);
 | 
|  |    650 |          A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |    651 |  \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs"
 | 
|  |    652 | apply (erule rev_mp)
 | 
|  |    653 | apply (erule rev_mp)
 | 
|  |    654 | apply (erule rev_mp)
 | 
|  |    655 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    656 | apply (frule_tac [5] Says_ticket_parts)
 | 
|  |    657 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    658 | apply (simp_all (no_asm_simp) add: all_conj_distrib)
 | 
|  |    659 | apply blast
 | 
|  |    660 | txt{*K3*}
 | 
|  |    661 | apply (blast dest: authK_authentic Says_Kas_message_form Says_Tgs_message_form)
 | 
|  |    662 | txt{*K4*}
 | 
|  |    663 | apply (force dest!: Crypt_imp_keysFor)
 | 
|  |    664 | txt{*K5*}
 | 
|  |    665 | apply (blast dest: Key_unique_SesKey)
 | 
|  |    666 | done
 | 
|  |    667 | 
 | 
|  |    668 | text{*Anticipated here from next subsection*}
 | 
|  |    669 | lemma unique_CryptKey:
 | 
|  |    670 |      "\<lbrakk> Crypt (shrK B)  \<lbrace>Agent A,  Agent B,  Key SesKey, T\<rbrace>
 | 
|  |    671 |            \<in> parts (spies evs);
 | 
|  |    672 |          Crypt (shrK B') \<lbrace>Agent A', Agent B', Key SesKey, T'\<rbrace>
 | 
|  |    673 |            \<in> parts (spies evs);  Key SesKey \<notin> analz (spies evs);
 | 
|  |    674 |          evs \<in> kerbV \<rbrakk>
 | 
|  |    675 |       \<Longrightarrow> A=A' & B=B' & T=T'"
 | 
|  |    676 | apply (erule rev_mp)
 | 
|  |    677 | apply (erule rev_mp)
 | 
|  |    678 | apply (erule rev_mp)
 | 
|  |    679 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    680 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    681 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    682 | txt{*Fake, K2, K4*}
 | 
|  |    683 | apply (blast+)
 | 
|  |    684 | done
 | 
|  |    685 | 
 | 
|  |    686 | lemma Says_K6:
 | 
|  |    687 |      "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
 | 
|  |    688 |          Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
 | 
|  |    689 |                       servTicket\<rbrace> \<in> set evs;
 | 
|  |    690 |          Key servK \<notin> analz (spies evs);
 | 
|  |    691 |          A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |    692 |       \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
 | 
|  |    693 | apply (frule Says_Tgs_message_form, assumption, clarify)
 | 
|  |    694 | apply (erule rev_mp)
 | 
|  |    695 | apply (erule rev_mp)
 | 
|  |    696 | apply (erule rev_mp)
 | 
|  |    697 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    698 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    699 | apply (frule_tac [5] Says_ticket_parts)
 | 
|  |    700 | apply (simp_all (no_asm_simp))
 | 
|  |    701 | 
 | 
|  |    702 | txt{*fake*}
 | 
|  |    703 | apply blast
 | 
|  |    704 | txt{*K4*}
 | 
|  |    705 | apply (force dest!: Crypt_imp_keysFor, clarify)
 | 
|  |    706 | txt{*K6*}
 | 
|  |    707 | apply (drule  Says_imp_spies [THEN parts.Inj, THEN parts.Fst])
 | 
|  |    708 | apply (drule  Says_imp_spies [THEN parts.Inj, THEN parts.Snd])
 | 
|  |    709 | apply (blast dest!: unique_CryptKey)
 | 
|  |    710 | done
 | 
|  |    711 | 
 | 
|  |    712 | text{*Needs a unicity theorem, hence moved here*}
 | 
|  |    713 | lemma servK_authentic_ter:
 | 
|  |    714 |  "\<lbrakk> Says Kas A
 | 
|  |    715 |        \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>, authTicket\<rbrace> \<in> set evs;
 | 
|  |    716 |      Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>
 | 
|  |    717 |        \<in> parts (spies evs);
 | 
|  |    718 |      Key authK \<notin> analz (spies evs);
 | 
|  |    719 |      evs \<in> kerbV \<rbrakk>
 | 
|  |    720 |  \<Longrightarrow> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Ts\<rbrace>, 
 | 
|  |    721 |                  Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace> \<rbrace>
 | 
|  |    722 |        \<in> set evs"
 | 
|  |    723 | apply (frule Says_Kas_message_form, assumption)
 | 
|  |    724 | apply clarify
 | 
|  |    725 | apply (erule rev_mp)
 | 
|  |    726 | apply (erule rev_mp)
 | 
|  |    727 | apply (erule rev_mp)
 | 
|  |    728 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    729 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    730 | apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
 | 
|  |    731 | txt{*K2 and K4 remain*}
 | 
|  |    732 | apply (blast dest!: servK_authentic Says_Tgs_message_form authKeys_used)
 | 
|  |    733 | apply (blast dest!: unique_CryptKey)
 | 
|  |    734 | done
 | 
|  |    735 | 
 | 
|  |    736 | 
 | 
|  |    737 | subsection{*Unicity Theorems*}
 | 
|  |    738 | 
 | 
|  |    739 | text{* The session key, if secure, uniquely identifies the Ticket
 | 
|  |    740 |    whether authTicket or servTicket. As a matter of fact, one can read
 | 
|  |    741 |    also Tgs in the place of B.                                     *}
 | 
|  |    742 | 
 | 
|  |    743 | 
 | 
|  |    744 | lemma unique_authKeys:
 | 
|  |    745 |      "\<lbrakk> Says Kas A
 | 
|  |    746 |               \<lbrace>Crypt Ka \<lbrace>Key authK, Agent Tgs, Ta\<rbrace>, X\<rbrace> \<in> set evs;
 | 
|  |    747 |          Says Kas A'
 | 
|  |    748 |               \<lbrace>Crypt Ka' \<lbrace>Key authK, Agent Tgs, Ta'\<rbrace>, X'\<rbrace> \<in> set evs;
 | 
|  |    749 |          evs \<in> kerbV \<rbrakk> \<Longrightarrow> A=A' \<and> Ka=Ka' \<and> Ta=Ta' \<and> X=X'"
 | 
|  |    750 | apply (erule rev_mp)
 | 
|  |    751 | apply (erule rev_mp)
 | 
|  |    752 | apply (erule kerbV.induct)
 | 
|  |    753 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    754 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    755 | apply blast+
 | 
|  |    756 | done
 | 
|  |    757 | 
 | 
|  |    758 | text{* servK uniquely identifies the message from Tgs *}
 | 
|  |    759 | lemma unique_servKeys:
 | 
|  |    760 |      "\<lbrakk> Says Tgs A
 | 
|  |    761 |               \<lbrace>Crypt K \<lbrace>Key servK, Agent B, Ts\<rbrace>, X\<rbrace> \<in> set evs;
 | 
|  |    762 |          Says Tgs A'
 | 
|  |    763 |               \<lbrace>Crypt K' \<lbrace>Key servK, Agent B', Ts'\<rbrace>, X'\<rbrace> \<in> set evs;
 | 
|  |    764 |          evs \<in> kerbV \<rbrakk> \<Longrightarrow> A=A' \<and> B=B' \<and> K=K' \<and> Ts=Ts' \<and> X=X'"
 | 
|  |    765 | apply (erule rev_mp)
 | 
|  |    766 | apply (erule rev_mp)
 | 
|  |    767 | apply (erule kerbV.induct)
 | 
|  |    768 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    769 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    770 | apply blast+
 | 
|  |    771 | done
 | 
|  |    772 | 
 | 
|  |    773 | subsection{*Lemmas About the Predicate @{term AKcryptSK}*}
 | 
|  |    774 | 
 | 
|  |    775 | lemma not_AKcryptSK_Nil [iff]: "\<not> AKcryptSK authK servK []"
 | 
|  |    776 | apply (simp add: AKcryptSK_def)
 | 
|  |    777 | done
 | 
|  |    778 | 
 | 
|  |    779 | lemma AKcryptSKI:
 | 
|  |    780 |  "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>, X \<rbrace> \<in> set evs;
 | 
|  |    781 |      evs \<in> kerbV \<rbrakk> \<Longrightarrow> AKcryptSK authK servK evs"
 | 
|  |    782 | apply (unfold AKcryptSK_def)
 | 
|  |    783 | apply (blast dest: Says_Tgs_message_form)
 | 
|  |    784 | done
 | 
|  |    785 | 
 | 
|  |    786 | lemma AKcryptSK_Says [simp]:
 | 
|  |    787 |    "AKcryptSK authK servK (Says S A X # evs) =
 | 
|  |    788 |      (S = Tgs \<and>
 | 
|  |    789 |       (\<exists>B tt. X = \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>,
 | 
|  |    790 |                     Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, tt\<rbrace> \<rbrace>)
 | 
|  |    791 |      | AKcryptSK authK servK evs)"
 | 
|  |    792 | apply (unfold AKcryptSK_def)
 | 
|  |    793 | apply (simp (no_asm))
 | 
|  |    794 | apply blast
 | 
|  |    795 | done
 | 
|  |    796 | 
 | 
|  |    797 | lemma AKcryptSK_Notes [simp]:
 | 
|  |    798 |    "AKcryptSK authK servK (Notes A X # evs) =
 | 
|  |    799 |       AKcryptSK authK servK evs"
 | 
|  |    800 | apply (unfold AKcryptSK_def)
 | 
|  |    801 | apply (simp (no_asm))
 | 
|  |    802 | done
 | 
|  |    803 | 
 | 
|  |    804 | (*A fresh authK cannot be associated with any other
 | 
|  |    805 |   (with respect to a given trace). *)
 | 
|  |    806 | lemma Auth_fresh_not_AKcryptSK:
 | 
|  |    807 |      "\<lbrakk> Key authK \<notin> used evs; evs \<in> kerbV \<rbrakk>
 | 
|  |    808 |       \<Longrightarrow> \<not> AKcryptSK authK servK evs"
 | 
|  |    809 | apply (unfold AKcryptSK_def)
 | 
|  |    810 | apply (erule rev_mp)
 | 
|  |    811 | apply (erule kerbV.induct)
 | 
|  |    812 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    813 | apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
 | 
|  |    814 | done
 | 
|  |    815 | 
 | 
|  |    816 | (*A fresh servK cannot be associated with any other
 | 
|  |    817 |   (with respect to a given trace). *)
 | 
|  |    818 | lemma Serv_fresh_not_AKcryptSK:
 | 
|  |    819 |  "Key servK \<notin> used evs \<Longrightarrow> \<not> AKcryptSK authK servK evs"
 | 
|  |    820 | apply (unfold AKcryptSK_def, blast)
 | 
|  |    821 | done
 | 
|  |    822 | 
 | 
|  |    823 | lemma authK_not_AKcryptSK:
 | 
|  |    824 |      "\<lbrakk> Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, tk\<rbrace>
 | 
|  |    825 |            \<in> parts (spies evs);  evs \<in> kerbV \<rbrakk>
 | 
|  |    826 |       \<Longrightarrow> \<not> AKcryptSK K authK evs"
 | 
|  |    827 | apply (erule rev_mp)
 | 
|  |    828 | apply (erule kerbV.induct)
 | 
|  |    829 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    830 | apply (frule_tac [5] Says_ticket_parts, simp_all)
 | 
|  |    831 | txt{*Fake*}
 | 
|  |    832 | apply blast
 | 
|  |    833 | txt{*K2: by freshness*}
 | 
|  |    834 | apply (simp add: AKcryptSK_def)
 | 
|  |    835 | apply blast
 | 
|  |    836 | txt{*K4*}
 | 
|  |    837 | apply blast
 | 
|  |    838 | done
 | 
|  |    839 | 
 | 
|  |    840 | text{*A secure serverkey cannot have been used to encrypt others*}
 | 
|  |    841 | lemma servK_not_AKcryptSK:
 | 
|  |    842 |  "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, tt\<rbrace> \<in> parts (spies evs);
 | 
|  |    843 |      Key SK \<notin> analz (spies evs);  SK \<in> symKeys;
 | 
|  |    844 |      B \<noteq> Tgs;  evs \<in> kerbV \<rbrakk>
 | 
|  |    845 |   \<Longrightarrow> \<not> AKcryptSK SK K evs"
 | 
|  |    846 | apply (erule rev_mp)
 | 
|  |    847 | apply (erule rev_mp)
 | 
|  |    848 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |    849 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    850 | apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
 | 
|  |    851 | txt{*K4 splits into distinct subcases*}
 | 
|  |    852 | apply auto
 | 
|  |    853 | txt{*servK can't have been enclosed in two certificates*}
 | 
|  |    854 |  prefer 2 apply (blast dest: unique_CryptKey)
 | 
|  |    855 | txt{*servK is fresh and so could not have been used, by
 | 
|  |    856 |    @{text new_keys_not_used}*}
 | 
|  |    857 | apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
 | 
|  |    858 | done
 | 
|  |    859 | 
 | 
|  |    860 | text{*Long term keys are not issued as servKeys*}
 | 
|  |    861 | lemma shrK_not_AKcryptSK:
 | 
|  |    862 |      "evs \<in> kerbV \<Longrightarrow> \<not> AKcryptSK K (shrK A) evs"
 | 
|  |    863 | apply (unfold AKcryptSK_def)
 | 
|  |    864 | apply (erule kerbV.induct)
 | 
|  |    865 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    866 | apply (frule_tac [5] Says_ticket_parts, auto)
 | 
|  |    867 | done
 | 
|  |    868 | 
 | 
|  |    869 | text{*The Tgs message associates servK with authK and therefore not with any
 | 
|  |    870 |   other key authK.*}
 | 
|  |    871 | lemma Says_Tgs_AKcryptSK:
 | 
|  |    872 |      "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>, X \<rbrace>
 | 
|  |    873 |            \<in> set evs;
 | 
|  |    874 |          authK' \<noteq> authK;  evs \<in> kerbV \<rbrakk>
 | 
|  |    875 |       \<Longrightarrow> \<not> AKcryptSK authK' servK evs"
 | 
|  |    876 | apply (unfold AKcryptSK_def)
 | 
|  |    877 | apply (blast dest: unique_servKeys)
 | 
|  |    878 | done
 | 
|  |    879 | 
 | 
|  |    880 | lemma AKcryptSK_not_AKcryptSK:
 | 
|  |    881 |      "\<lbrakk> AKcryptSK authK servK evs;  evs \<in> kerbV \<rbrakk>
 | 
|  |    882 |       \<Longrightarrow> \<not> AKcryptSK servK K evs"
 | 
|  |    883 | apply (erule rev_mp)
 | 
|  |    884 | apply (erule kerbV.induct)
 | 
|  |    885 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    886 | apply (frule_tac [5] Says_ticket_parts)
 | 
|  |    887 | apply (simp_all, safe)
 | 
|  |    888 | txt{*K4 splits into subcases*}
 | 
|  |    889 | (*apply simp_all*)
 | 
|  |    890 | prefer 4 apply (blast dest!: authK_not_AKcryptSK)
 | 
|  |    891 | txt{*servK is fresh and so could not have been used, by
 | 
|  |    892 |    @{text new_keys_not_used}*}
 | 
|  |    893 |  prefer 2 
 | 
|  |    894 |  apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
 | 
|  |    895 | txt{*Others by freshness*}
 | 
|  |    896 | apply (blast+)
 | 
|  |    897 | done
 | 
|  |    898 | 
 | 
|  |    899 | lemma not_different_AKcryptSK:
 | 
|  |    900 |      "\<lbrakk> AKcryptSK authK servK evs;
 | 
|  |    901 |         authK' \<noteq> authK;  evs \<in> kerbV \<rbrakk>
 | 
|  |    902 |       \<Longrightarrow> \<not> AKcryptSK authK' servK evs  \<and> servK \<in> symKeys"
 | 
|  |    903 | apply (simp add: AKcryptSK_def)
 | 
|  |    904 | apply (blast dest: unique_servKeys Says_Tgs_message_form)
 | 
|  |    905 | done
 | 
|  |    906 | 
 | 
|  |    907 | lemma AKcryptSK_not_AKcryptSK:
 | 
|  |    908 |      "\<lbrakk> AKcryptSK authK servK evs;  evs \<in> kerbV \<rbrakk>
 | 
|  |    909 |       \<Longrightarrow> \<not> AKcryptSK servK K evs"
 | 
|  |    910 | apply (erule rev_mp)
 | 
|  |    911 | apply (erule kerbV.induct)
 | 
|  |    912 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |    913 | apply (frule_tac [5] Says_ticket_parts, simp_all, safe)
 | 
|  |    914 | txt{*K4 splits into subcases*}
 | 
|  |    915 | apply simp_all
 | 
|  |    916 | prefer 4 apply (blast dest!: authK_not_AKcryptSK)
 | 
|  |    917 | txt{*servK is fresh and so could not have been used, by
 | 
|  |    918 |    @{text new_keys_not_used}*}
 | 
|  |    919 |  prefer 2 
 | 
|  |    920 |  apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
 | 
|  |    921 | txt{*Others by freshness*}
 | 
|  |    922 | apply (blast+)
 | 
|  |    923 | done
 | 
|  |    924 | 
 | 
|  |    925 | text{*The only session keys that can be found with the help of session keys are
 | 
|  |    926 |   those sent by Tgs in step K4.  *}
 | 
|  |    927 | 
 | 
|  |    928 | text{*We take some pains to express the property
 | 
|  |    929 |   as a logical equivalence so that the simplifier can apply it.*}
 | 
|  |    930 | lemma Key_analz_image_Key_lemma:
 | 
|  |    931 |      "P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) \<longrightarrow> (K:KK | Key K \<in> analz H)
 | 
|  |    932 |       \<Longrightarrow>
 | 
|  |    933 |       P \<longrightarrow> (Key K \<in> analz (Key`KK Un H)) = (K:KK | Key K \<in> analz H)"
 | 
|  |    934 | by (blast intro: analz_mono [THEN subsetD])
 | 
|  |    935 | 
 | 
|  |    936 | 
 | 
|  |    937 | lemma AKcryptSK_analz_insert:
 | 
|  |    938 |      "\<lbrakk> AKcryptSK K K' evs; K \<in> symKeys; evs \<in> kerbV \<rbrakk>
 | 
|  |    939 |       \<Longrightarrow> Key K' \<in> analz (insert (Key K) (spies evs))"
 | 
|  |    940 | apply (simp add: AKcryptSK_def, clarify)
 | 
|  |    941 | apply (drule Says_imp_spies [THEN analz.Inj, THEN analz_insertI], auto)
 | 
|  |    942 | done
 | 
|  |    943 | 
 | 
|  |    944 | lemma authKeys_are_not_AKcryptSK:
 | 
|  |    945 |      "\<lbrakk> K \<in> authKeys evs Un range shrK;  evs \<in> kerbV \<rbrakk>
 | 
|  |    946 |       \<Longrightarrow> \<forall>SK. \<not> AKcryptSK SK K evs \<and> K \<in> symKeys"
 | 
|  |    947 | apply (simp add: authKeys_def AKcryptSK_def)
 | 
|  |    948 | apply (blast dest: Says_Kas_message_form Says_Tgs_message_form)
 | 
|  |    949 | done
 | 
|  |    950 | 
 | 
|  |    951 | lemma not_authKeys_not_AKcryptSK:
 | 
|  |    952 |      "\<lbrakk> K \<notin> authKeys evs;
 | 
|  |    953 |          K \<notin> range shrK; evs \<in> kerbV \<rbrakk>
 | 
|  |    954 |       \<Longrightarrow> \<forall>SK. \<not> AKcryptSK K SK evs"
 | 
|  |    955 | apply (simp add: AKcryptSK_def)
 | 
|  |    956 | apply (blast dest: Says_Tgs_message_form)
 | 
|  |    957 | done
 | 
|  |    958 | 
 | 
|  |    959 | 
 | 
|  |    960 | subsection{*Secrecy Theorems*}
 | 
|  |    961 | 
 | 
|  |    962 | text{*For the Oops2 case of the next theorem*}
 | 
|  |    963 | lemma Oops2_not_AKcryptSK:
 | 
|  |    964 |      "\<lbrakk> evs \<in> kerbV;
 | 
|  |    965 |          Says Tgs A \<lbrace>Crypt authK
 | 
|  |    966 |                      \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |    967 |            \<in> set evs \<rbrakk>
 | 
|  |    968 |       \<Longrightarrow> \<not> AKcryptSK servK SK evs"
 | 
|  |    969 | apply (blast dest: AKcryptSKI AKcryptSK_not_AKcryptSK)
 | 
|  |    970 | done
 | 
|  |    971 |    
 | 
|  |    972 | text{* Big simplification law for keys SK that are not crypted by keys in KK
 | 
|  |    973 |  It helps prove three, otherwise hard, facts about keys. These facts are
 | 
|  |    974 |  exploited as simplification laws for analz, and also "limit the damage"
 | 
|  |    975 |  in case of loss of a key to the spy. See ESORICS98.*}
 | 
|  |    976 | lemma Key_analz_image_Key [rule_format (no_asm)]:
 | 
|  |    977 |      "evs \<in> kerbV \<Longrightarrow>
 | 
|  |    978 |       (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) \<longrightarrow>
 | 
|  |    979 |        (\<forall>K \<in> KK. \<not> AKcryptSK K SK evs)   \<longrightarrow>
 | 
|  |    980 |        (Key SK \<in> analz (Key`KK Un (spies evs))) =
 | 
|  |    981 |        (SK \<in> KK | Key SK \<in> analz (spies evs)))"
 | 
|  |    982 | apply (erule kerbV.induct)
 | 
|  |    983 | apply (frule_tac [10] Oops_range_spies2)
 | 
|  |    984 | apply (frule_tac [9] Oops_range_spies1)
 | 
|  |    985 | (*Used to apply Says_tgs_message form, which is no longer available. 
 | 
|  |    986 |   Instead\<dots>*)
 | 
|  |    987 | apply (drule_tac [7] Says_ticket_analz)
 | 
|  |    988 | (*Used to apply Says_kas_message form, which is no longer available. 
 | 
|  |    989 |   Instead\<dots>*)
 | 
|  |    990 | apply (drule_tac [5] Says_ticket_analz)
 | 
|  |    991 | apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
 | 
|  |    992 | txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
 | 
|  |    993 |  the induction hypothesis*}
 | 
|  |    994 | apply (case_tac [9] "AKcryptSK authK SK evsO1")
 | 
|  |    995 | apply (case_tac [7] "AKcryptSK servK SK evs5")
 | 
|  |    996 | apply (simp_all del: image_insert
 | 
|  |    997 |           add: analz_image_freshK_simps AKcryptSK_Says shrK_not_AKcryptSK
 | 
|  |    998 |                Oops2_not_AKcryptSK Auth_fresh_not_AKcryptSK
 | 
|  |    999 |                Serv_fresh_not_AKcryptSK Says_Tgs_AKcryptSK Spy_analz_shrK)
 | 
|  |   1000 | txt{*Fake*} 
 | 
|  |   1001 | apply spy_analz
 | 
|  |   1002 | txt{*K2*}
 | 
|  |   1003 | apply blast 
 | 
|  |   1004 | txt{*Cases K3 and K5 solved by the simplifier thanks to the ticket being in 
 | 
|  |   1005 | analz - this strategy is new wrt version IV*} 
 | 
|  |   1006 | txt{*K4*}
 | 
|  |   1007 | apply (blast dest!: authK_not_AKcryptSK)
 | 
|  |   1008 | txt{*Oops1*}
 | 
|  |   1009 | apply clarify
 | 
|  |   1010 | apply simp
 | 
|  |   1011 | apply (blast dest!: AKcryptSK_analz_insert)
 | 
|  |   1012 | done
 | 
|  |   1013 | 
 | 
|  |   1014 | text{* First simplification law for analz: no session keys encrypt
 | 
|  |   1015 | authentication keys or shared keys. *}
 | 
|  |   1016 | lemma analz_insert_freshK1:
 | 
|  |   1017 |      "\<lbrakk> evs \<in> kerbV;  K \<in> authKeys evs Un range shrK;
 | 
|  |   1018 |         SesKey \<notin> range shrK \<rbrakk>
 | 
|  |   1019 |       \<Longrightarrow> (Key K \<in> analz (insert (Key SesKey) (spies evs))) =
 | 
|  |   1020 |           (K = SesKey | Key K \<in> analz (spies evs))"
 | 
|  |   1021 | apply (frule authKeys_are_not_AKcryptSK, assumption)
 | 
|  |   1022 | apply (simp del: image_insert
 | 
|  |   1023 |             add: analz_image_freshK_simps add: Key_analz_image_Key)
 | 
|  |   1024 | done
 | 
|  |   1025 | 
 | 
|  |   1026 | 
 | 
|  |   1027 | text{* Second simplification law for analz: no service keys encrypt any other keys.*}
 | 
|  |   1028 | lemma analz_insert_freshK2:
 | 
|  |   1029 |      "\<lbrakk> evs \<in> kerbV;  servK \<notin> (authKeys evs); servK \<notin> range shrK;
 | 
|  |   1030 |         K \<in> symKeys \<rbrakk>
 | 
|  |   1031 |       \<Longrightarrow> (Key K \<in> analz (insert (Key servK) (spies evs))) =
 | 
|  |   1032 |           (K = servK | Key K \<in> analz (spies evs))"
 | 
|  |   1033 | apply (frule not_authKeys_not_AKcryptSK, assumption, assumption)
 | 
|  |   1034 | apply (simp del: image_insert
 | 
|  |   1035 |             add: analz_image_freshK_simps add: Key_analz_image_Key)
 | 
|  |   1036 | done
 | 
|  |   1037 | 
 | 
|  |   1038 | 
 | 
|  |   1039 | text{* Third simplification law for analz: only one authentication key encrypts a certain service key.*}
 | 
|  |   1040 | 
 | 
|  |   1041 | lemma analz_insert_freshK3:
 | 
|  |   1042 |  "\<lbrakk> AKcryptSK authK servK evs;
 | 
|  |   1043 |     authK' \<noteq> authK; authK' \<notin> range shrK; evs \<in> kerbV \<rbrakk>
 | 
|  |   1044 |         \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
 | 
|  |   1045 |                 (servK = authK' | Key servK \<in> analz (spies evs))"
 | 
|  |   1046 | apply (drule_tac authK' = authK' in not_different_AKcryptSK, blast, assumption)
 | 
|  |   1047 | apply (simp del: image_insert
 | 
|  |   1048 |             add: analz_image_freshK_simps add: Key_analz_image_Key)
 | 
|  |   1049 | done
 | 
|  |   1050 | 
 | 
|  |   1051 | lemma analz_insert_freshK3_bis:
 | 
|  |   1052 |  "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |   1053 |         \<in> set evs; 
 | 
|  |   1054 |      authK \<noteq> authK'; authK' \<notin> range shrK; evs \<in> kerbV \<rbrakk>
 | 
|  |   1055 |         \<Longrightarrow> (Key servK \<in> analz (insert (Key authK') (spies evs))) =
 | 
|  |   1056 |                 (servK = authK' | Key servK \<in> analz (spies evs))"
 | 
|  |   1057 | apply (frule AKcryptSKI, assumption)
 | 
|  |   1058 | apply (simp add: analz_insert_freshK3)
 | 
|  |   1059 | done
 | 
|  |   1060 | 
 | 
|  |   1061 | text{*a weakness of the protocol*}
 | 
|  |   1062 | lemma authK_compromises_servK:
 | 
|  |   1063 |      "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |   1064 |         \<in> set evs;  authK \<in> symKeys;
 | 
|  |   1065 |          Key authK \<in> analz (spies evs); evs \<in> kerbV \<rbrakk>
 | 
|  |   1066 |       \<Longrightarrow> Key servK \<in> analz (spies evs)"
 | 
|  |   1067 | apply (force dest: Says_imp_spies [THEN analz.Inj, THEN analz.Fst, THEN analz.Decrypt, THEN analz.Fst])
 | 
|  |   1068 | done
 | 
|  |   1069 | 
 | 
|  |   1070 | text{*lemma @{text servK_notin_authKeysD} not needed in version V*}
 | 
|  |   1071 | 
 | 
|  |   1072 | text{*If Spy sees the Authentication Key sent in msg K2, then
 | 
|  |   1073 |     the Key has expired.*}
 | 
|  |   1074 | lemma Confidentiality_Kas_lemma [rule_format]:
 | 
|  |   1075 |      "\<lbrakk> authK \<in> symKeys; A \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |   1076 |       \<Longrightarrow> Says Kas A
 | 
|  |   1077 |                \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
 | 
|  |   1078 |           Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK, Number Ta\<rbrace>\<rbrace>
 | 
|  |   1079 |             \<in> set evs \<longrightarrow>
 | 
|  |   1080 |           Key authK \<in> analz (spies evs) \<longrightarrow>
 | 
|  |   1081 |           expiredAK Ta evs"
 | 
|  |   1082 | apply (erule kerbV.induct)
 | 
|  |   1083 | apply (frule_tac [10] Oops_range_spies2)
 | 
|  |   1084 | apply (frule_tac [9] Oops_range_spies1)
 | 
|  |   1085 | apply (frule_tac [7] Says_ticket_analz)
 | 
|  |   1086 | apply (frule_tac [5] Says_ticket_analz)
 | 
|  |   1087 | apply (safe del: impI conjI impCE)
 | 
|  |   1088 | apply (simp_all (no_asm_simp) add: Says_Kas_message_form less_SucI analz_insert_eq not_parts_not_analz analz_insert_freshK1 pushes)
 | 
|  |   1089 | txt{*Fake*}
 | 
|  |   1090 | apply spy_analz
 | 
|  |   1091 | txt{*K2*}
 | 
|  |   1092 | apply blast
 | 
|  |   1093 | txt{*K4*}
 | 
|  |   1094 | apply blast
 | 
|  |   1095 | txt{*Oops1*}
 | 
|  |   1096 | apply (blast dest!: unique_authKeys intro: less_SucI)
 | 
|  |   1097 | txt{*Oops2*}
 | 
|  |   1098 | apply (blast dest: Says_Tgs_message_form Says_Kas_message_form)
 | 
|  |   1099 | done
 | 
|  |   1100 | 
 | 
|  |   1101 | lemma Confidentiality_Kas:
 | 
|  |   1102 |      "\<lbrakk> Says Kas A
 | 
|  |   1103 |               \<lbrace>Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>, authTicket\<rbrace>
 | 
|  |   1104 |            \<in> set evs;
 | 
|  |   1105 |         \<not> expiredAK Ta evs;
 | 
|  |   1106 |         A \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |   1107 |       \<Longrightarrow> Key authK \<notin> analz (spies evs)"
 | 
|  |   1108 | apply (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma)
 | 
|  |   1109 | done
 | 
|  |   1110 | 
 | 
|  |   1111 | text{*If Spy sees the Service Key sent in msg K4, then
 | 
|  |   1112 |     the Key has expired.*}
 | 
|  |   1113 | 
 | 
|  |   1114 | lemma Confidentiality_lemma [rule_format]:
 | 
|  |   1115 |      "\<lbrakk> Says Tgs A
 | 
|  |   1116 | 	    \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
 | 
|  |   1117 | 	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
 | 
|  |   1118 | 	   \<in> set evs;
 | 
|  |   1119 | 	Key authK \<notin> analz (spies evs);
 | 
|  |   1120 |         servK \<in> symKeys;
 | 
|  |   1121 | 	A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1122 |       \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
 | 
|  |   1123 | 	  expiredSK Ts evs"
 | 
|  |   1124 | apply (erule rev_mp)
 | 
|  |   1125 | apply (erule rev_mp)
 | 
|  |   1126 | apply (erule kerbV.induct)
 | 
|  |   1127 | apply (rule_tac [9] impI)+;
 | 
|  |   1128 |   --{*The Oops1 case is unusual: must simplify
 | 
|  |   1129 |     @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
 | 
|  |   1130 |    @{text analz_mono_contra} weaken it to
 | 
|  |   1131 |    @{term "Authkey \<notin> analz (spies evs)"},
 | 
|  |   1132 |   for we then conclude @{term "authK \<noteq> authKa"}.*}
 | 
|  |   1133 | apply analz_mono_contra
 | 
|  |   1134 | apply (frule_tac [10] Oops_range_spies2)
 | 
|  |   1135 | apply (frule_tac [9] Oops_range_spies1)
 | 
|  |   1136 | apply (frule_tac [7] Says_ticket_analz)
 | 
|  |   1137 | apply (frule_tac [5] Says_ticket_analz)
 | 
|  |   1138 | apply (safe del: impI conjI impCE)
 | 
|  |   1139 | apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
 | 
|  |   1140 | txt{*Fake*}
 | 
|  |   1141 | apply spy_analz
 | 
|  |   1142 | txt{*K2*}
 | 
|  |   1143 | apply (blast intro: parts_insertI less_SucI)
 | 
|  |   1144 | txt{*K4*}
 | 
|  |   1145 | apply (blast dest: authTicket_authentic Confidentiality_Kas)
 | 
|  |   1146 | txt{*Oops1*}
 | 
|  |   1147 |  apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
 | 
|  |   1148 | txt{*Oops2*}
 | 
|  |   1149 |   apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
 | 
|  |   1150 | done
 | 
|  |   1151 | 
 | 
|  |   1152 | 
 | 
|  |   1153 | text{* In the real world Tgs can't check wheter authK is secure! *}
 | 
|  |   1154 | lemma Confidentiality_Tgs:
 | 
|  |   1155 |      "\<lbrakk> Says Tgs A
 | 
|  |   1156 |               \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |   1157 |            \<in> set evs;
 | 
|  |   1158 |          Key authK \<notin> analz (spies evs);
 | 
|  |   1159 |          \<not> expiredSK Ts evs;
 | 
|  |   1160 |          A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1161 |       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 | 
|  |   1162 | apply (blast dest: Says_Tgs_message_form Confidentiality_lemma)
 | 
|  |   1163 | done
 | 
|  |   1164 | 
 | 
|  |   1165 | text{* In the real world Tgs CAN check what Kas sends! *}
 | 
|  |   1166 | lemma Confidentiality_Tgs_bis:
 | 
|  |   1167 |      "\<lbrakk> Says Kas A
 | 
|  |   1168 |                \<lbrace>Crypt Ka \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>, authTicket\<rbrace>
 | 
|  |   1169 |            \<in> set evs;
 | 
|  |   1170 |          Says Tgs A
 | 
|  |   1171 |               \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |   1172 |            \<in> set evs;
 | 
|  |   1173 |          \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
 | 
|  |   1174 |          A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1175 |       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 | 
|  |   1176 | apply (blast dest!: Confidentiality_Kas Confidentiality_Tgs)
 | 
|  |   1177 | done
 | 
|  |   1178 | 
 | 
|  |   1179 | text{*Most general form*}
 | 
|  |   1180 | lemmas Confidentiality_Tgs_ter = authTicket_authentic [THEN Confidentiality_Tgs_bis]
 | 
|  |   1181 | 
 | 
|  |   1182 | lemmas Confidentiality_Auth_A = authK_authentic [THEN exE, THEN Confidentiality_Kas]
 | 
|  |   1183 | 
 | 
|  |   1184 | text{*Needs a confidentiality guarantee, hence moved here.
 | 
|  |   1185 |       Authenticity of servK for A*}
 | 
|  |   1186 | lemma servK_authentic_bis_r:
 | 
|  |   1187 |      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
 | 
|  |   1188 |            \<in> parts (spies evs);
 | 
|  |   1189 |          Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1190 |            \<in> parts (spies evs);
 | 
|  |   1191 |          \<not> expiredAK Ta evs; A \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1192 |  \<Longrightarrow> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, 
 | 
|  |   1193 |                  Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>
 | 
|  |   1194 |        \<in> set evs"
 | 
|  |   1195 | apply (frule authK_authentic, assumption, assumption)
 | 
|  |   1196 | apply (erule exE)
 | 
|  |   1197 | apply (drule Confidentiality_Auth_A, assumption, assumption)
 | 
|  |   1198 | apply (blast, assumption, assumption, assumption)
 | 
|  |   1199 | apply (blast dest:  servK_authentic_ter)
 | 
|  |   1200 | done
 | 
|  |   1201 | 
 | 
|  |   1202 | lemma Confidentiality_Serv_A:
 | 
|  |   1203 |      "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
 | 
|  |   1204 |            \<in> parts (spies evs);
 | 
|  |   1205 |          Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1206 |            \<in> parts (spies evs);
 | 
|  |   1207 |          \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
 | 
|  |   1208 |          A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
 | 
|  |   1209 |       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 | 
|  |   1210 | apply (drule authK_authentic, assumption, assumption)
 | 
|  |   1211 | apply (blast dest: Confidentiality_Kas Says_Kas_message_form servK_authentic_ter Confidentiality_Tgs_bis)
 | 
|  |   1212 | done
 | 
|  |   1213 | 
 | 
|  |   1214 | lemma Confidentiality_B:
 | 
|  |   1215 |      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |   1216 |            \<in> parts (spies evs);
 | 
|  |   1217 |          Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1218 |            \<in> parts (spies evs);
 | 
|  |   1219 |          Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
 | 
|  |   1220 |            \<in> parts (spies evs);
 | 
|  |   1221 |          \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
 | 
|  |   1222 |          A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
 | 
|  |   1223 |       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 | 
|  |   1224 | apply (frule authK_authentic)
 | 
|  |   1225 | apply (erule_tac [3] exE)
 | 
|  |   1226 | apply (frule_tac [3] Confidentiality_Kas)
 | 
|  |   1227 | apply (frule_tac [6] servTicket_authentic, auto)
 | 
|  |   1228 | apply (blast dest!: Confidentiality_Tgs_bis dest: Says_Kas_message_form servK_authentic unique_servKeys unique_authKeys)
 | 
|  |   1229 | done
 | 
|  |   1230 | 
 | 
|  |   1231 | lemma u_Confidentiality_B:
 | 
|  |   1232 |      "\<lbrakk> Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |   1233 |            \<in> parts (spies evs);
 | 
|  |   1234 |          \<not> expiredSK Ts evs;
 | 
|  |   1235 |          A \<notin> bad;  B \<notin> bad;  B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
 | 
|  |   1236 |       \<Longrightarrow> Key servK \<notin> analz (spies evs)"
 | 
|  |   1237 | apply (blast dest: u_servTicket_authentic u_NotexpiredSK_NotexpiredAK Confidentiality_Tgs_bis)
 | 
|  |   1238 | done
 | 
|  |   1239 | 
 | 
|  |   1240 | 
 | 
|  |   1241 | 
 | 
|  |   1242 | subsection{*Parties authentication: each party verifies "the identity of
 | 
|  |   1243 |        another party who generated some data" (quoted from Neuman and Ts'o).*}
 | 
|  |   1244 | 
 | 
|  |   1245 | text{*These guarantees don't assess whether two parties agree on
 | 
|  |   1246 |       the same session key: sending a message containing a key
 | 
|  |   1247 |       doesn't a priori state knowledge of the key.*}
 | 
|  |   1248 | 
 | 
|  |   1249 | 
 | 
|  |   1250 | text{*These didn't have existential form in version IV*}
 | 
|  |   1251 | lemma B_authenticates_A:
 | 
|  |   1252 |      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
 | 
|  |   1253 |         Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |   1254 |            \<in> parts (spies evs);
 | 
|  |   1255 |         Key servK \<notin> analz (spies evs);
 | 
|  |   1256 |         A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
 | 
|  |   1257 |   \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
 | 
|  |   1258 | apply (blast dest: servTicket_authentic_Tgs intro: Says_K5)
 | 
|  |   1259 | done
 | 
|  |   1260 | 
 | 
|  |   1261 | text{*The second assumption tells B what kind of key servK is.*}
 | 
|  |   1262 | lemma B_authenticates_A_r:
 | 
|  |   1263 |      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
 | 
|  |   1264 |          Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |   1265 |            \<in> parts (spies evs);
 | 
|  |   1266 |          Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1267 |            \<in> parts (spies evs);
 | 
|  |   1268 |          Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
 | 
|  |   1269 |            \<in> parts (spies evs);
 | 
|  |   1270 |          \<not> expiredSK Ts evs; \<not> expiredAK Ta evs;
 | 
|  |   1271 |          B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |   1272 |   \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
 | 
|  |   1273 | apply (blast intro: Says_K5 dest: Confidentiality_B servTicket_authentic_Tgs)
 | 
|  |   1274 | done
 | 
|  |   1275 | 
 | 
|  |   1276 | text{* @{text u_B_authenticates_A} would be the same as @{text B_authenticates_A} because the
 | 
|  |   1277 |  servK confidentiality assumption is yet unrelaxed*}
 | 
|  |   1278 | 
 | 
|  |   1279 | lemma u_B_authenticates_A_r:
 | 
|  |   1280 |      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
 | 
|  |   1281 |          Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |   1282 |            \<in> parts (spies evs);
 | 
|  |   1283 |          \<not> expiredSK Ts evs;
 | 
|  |   1284 |          B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |   1285 |   \<Longrightarrow> \<exists> ST. Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<rbrace> \<in> set evs"
 | 
|  |   1286 | apply (blast intro: Says_K5 dest: u_Confidentiality_B servTicket_authentic_Tgs)
 | 
|  |   1287 | done
 | 
|  |   1288 | 
 | 
|  |   1289 | lemma A_authenticates_B:
 | 
|  |   1290 |      "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
 | 
|  |   1291 |          Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1292 |            \<in> parts (spies evs);
 | 
|  |   1293 |          Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
 | 
|  |   1294 |            \<in> parts (spies evs);
 | 
|  |   1295 |          Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
 | 
|  |   1296 |          A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1297 |       \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
 | 
|  |   1298 | apply (frule authK_authentic)
 | 
|  |   1299 | apply assumption+
 | 
|  |   1300 | apply (frule servK_authentic)
 | 
|  |   1301 | prefer 2 apply (blast dest: authK_authentic Says_Kas_message_form)
 | 
|  |   1302 | apply assumption+
 | 
|  |   1303 | apply clarify
 | 
|  |   1304 | apply (blast dest: K4_imp_K2 Key_unique_SesKey intro!: Says_K6)
 | 
|  |   1305 | (*Single command proof: much slower!
 | 
|  |   1306 | apply (blast dest: authK_authentic servK_authentic Says_Kas_message_form Key_unique_SesKey K4_imp_K2 intro!: Says_K6)
 | 
|  |   1307 | *)
 | 
|  |   1308 | done
 | 
|  |   1309 | 
 | 
|  |   1310 | lemma A_authenticates_B_r:
 | 
|  |   1311 |      "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
 | 
|  |   1312 |          Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1313 |            \<in> parts (spies evs);
 | 
|  |   1314 |          Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
 | 
|  |   1315 |            \<in> parts (spies evs);
 | 
|  |   1316 |          \<not> expiredAK Ta evs; \<not> expiredSK Ts evs;
 | 
|  |   1317 |          A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1318 |       \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
 | 
|  |   1319 | apply (frule authK_authentic)
 | 
|  |   1320 | apply (erule_tac [3] exE)
 | 
|  |   1321 | apply (frule_tac [3] Says_Kas_message_form)
 | 
|  |   1322 | apply (frule_tac [4] Confidentiality_Kas)
 | 
|  |   1323 | apply (frule_tac [7] servK_authentic)
 | 
|  |   1324 | prefer 8 apply blast
 | 
|  |   1325 | apply (erule_tac [9] exE)
 | 
|  |   1326 | apply (erule_tac [9] exE)
 | 
|  |   1327 | apply (frule_tac [9] K4_imp_K2)
 | 
|  |   1328 | apply assumption+
 | 
|  |   1329 | apply (blast dest: Key_unique_SesKey intro!: Says_K6 dest: Confidentiality_Tgs
 | 
|  |   1330 | )
 | 
|  |   1331 | done
 | 
|  |   1332 | 
 | 
|  |   1333 | 
 | 
|  |   1334 | 
 | 
|  |   1335 | 
 | 
|  |   1336 | subsection{*Parties' knowledge of session keys. 
 | 
|  |   1337 |        An agent knows a session key if he used it to issue a cipher. These
 | 
|  |   1338 |        guarantees can be interpreted both in terms of key distribution
 | 
|  |   1339 |        and of non-injective agreement on the session key.*}
 | 
|  |   1340 | 
 | 
|  |   1341 | lemma Kas_Issues_A:
 | 
|  |   1342 |    "\<lbrakk> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>, authTicket\<rbrace> \<in> set evs;
 | 
|  |   1343 |       evs \<in> kerbV \<rbrakk>
 | 
|  |   1344 |   \<Longrightarrow> Kas Issues A with (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>)
 | 
|  |   1345 |           on evs"
 | 
|  |   1346 | apply (simp (no_asm) add: Issues_def)
 | 
|  |   1347 | apply (rule exI)
 | 
|  |   1348 | apply (rule conjI, assumption)
 | 
|  |   1349 | apply (simp (no_asm))
 | 
|  |   1350 | apply (erule rev_mp)
 | 
|  |   1351 | apply (erule kerbV.induct)
 | 
|  |   1352 | apply (frule_tac [5] Says_ticket_parts)
 | 
|  |   1353 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |   1354 | apply (simp_all (no_asm_simp) add: all_conj_distrib)
 | 
|  |   1355 | txt{*K2*}
 | 
|  |   1356 | apply (simp add: takeWhile_tail)
 | 
|  |   1357 | apply (blast dest: authK_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD])
 | 
|  |   1358 | done
 | 
|  |   1359 | 
 | 
|  |   1360 | lemma A_authenticates_and_keydist_to_Kas:
 | 
|  |   1361 |   "\<lbrakk> Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace> \<in> parts (spies evs);
 | 
|  |   1362 |      A \<notin> bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1363 |  \<Longrightarrow> Kas Issues A with (Crypt (shrK A) \<lbrace>Key authK, Peer, Ta\<rbrace>) 
 | 
|  |   1364 |           on evs"
 | 
|  |   1365 | by (blast dest!: authK_authentic Kas_Issues_A)
 | 
|  |   1366 | 
 | 
|  |   1367 | lemma Tgs_Issues_A:
 | 
|  |   1368 |     "\<lbrakk> Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>, servTicket\<rbrace>
 | 
|  |   1369 |          \<in> set evs; 
 | 
|  |   1370 |        Key authK \<notin> analz (spies evs);  evs \<in> kerbV \<rbrakk>
 | 
|  |   1371 |   \<Longrightarrow> Tgs Issues A with 
 | 
|  |   1372 |           (Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>) on evs"
 | 
|  |   1373 | apply (simp (no_asm) add: Issues_def)
 | 
|  |   1374 | apply (rule exI)
 | 
|  |   1375 | apply (rule conjI, assumption)
 | 
|  |   1376 | apply (simp (no_asm))
 | 
|  |   1377 | apply (erule rev_mp)
 | 
|  |   1378 | apply (erule rev_mp)
 | 
|  |   1379 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |   1380 | apply (frule_tac [5] Says_ticket_parts)
 | 
|  |   1381 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |   1382 | apply (simp_all (no_asm_simp) add: all_conj_distrib)
 | 
|  |   1383 | apply (simp add: takeWhile_tail)
 | 
|  |   1384 | (*Last two thms installed only to derive authK \<notin> range shrK*)
 | 
|  |   1385 | apply (blast dest: servK_authentic parts_spies_takeWhile_mono [THEN subsetD]
 | 
|  |   1386 |       parts_spies_evs_revD2 [THEN subsetD] authTicket_authentic 
 | 
|  |   1387 |       Says_Kas_message_form)
 | 
|  |   1388 | done
 | 
|  |   1389 | 
 | 
|  |   1390 | lemma A_authenticates_and_keydist_to_Tgs:
 | 
|  |   1391 |      "\<lbrakk> Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1392 |            \<in> parts (spies evs);
 | 
|  |   1393 |        Key authK \<notin> analz (spies evs); B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
 | 
|  |   1394 |   \<Longrightarrow> \<exists>A. Tgs Issues A with 
 | 
|  |   1395 |           (Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>) on evs"
 | 
|  |   1396 | by (blast dest: Tgs_Issues_A servK_authentic_bis)
 | 
|  |   1397 | 
 | 
|  |   1398 | lemma B_Issues_A:
 | 
|  |   1399 |      "\<lbrakk> Says B A (Crypt servK (Number T3)) \<in> set evs;
 | 
|  |   1400 |          Key servK \<notin> analz (spies evs);
 | 
|  |   1401 |          A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
 | 
|  |   1402 |       \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
 | 
|  |   1403 | apply (simp (no_asm) add: Issues_def)
 | 
|  |   1404 | apply (rule exI)
 | 
|  |   1405 | apply (rule conjI, assumption)
 | 
|  |   1406 | apply (simp (no_asm))
 | 
|  |   1407 | apply (erule rev_mp)
 | 
|  |   1408 | apply (erule rev_mp)
 | 
|  |   1409 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |   1410 | apply (simp_all (no_asm_simp) add: all_conj_distrib)
 | 
|  |   1411 | apply blast
 | 
|  |   1412 | txt{*K6 requires numerous lemmas*}
 | 
|  |   1413 | apply (simp add: takeWhile_tail)
 | 
|  |   1414 | apply (blast intro: Says_K6 dest: servTicket_authentic 
 | 
|  |   1415 |         parts_spies_takeWhile_mono [THEN subsetD] 
 | 
|  |   1416 |         parts_spies_evs_revD2 [THEN subsetD])
 | 
|  |   1417 | done
 | 
|  |   1418 | 
 | 
|  |   1419 | lemma A_authenticates_and_keydist_to_B:
 | 
|  |   1420 |      "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
 | 
|  |   1421 |          Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>
 | 
|  |   1422 |            \<in> parts (spies evs);
 | 
|  |   1423 |          Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>
 | 
|  |   1424 |            \<in> parts (spies evs);
 | 
|  |   1425 |          Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
 | 
|  |   1426 |          A \<notin> bad;  B \<notin> bad; B \<noteq> Tgs; evs \<in> kerbV \<rbrakk>
 | 
|  |   1427 |       \<Longrightarrow> B Issues A with (Crypt servK (Number T3)) on evs"
 | 
|  |   1428 | by (blast dest!: A_authenticates_B B_Issues_A)
 | 
|  |   1429 | 
 | 
|  |   1430 | 
 | 
|  |   1431 | (*Must use \<le> rather than =, otherwise it cannot be proved inductively!*)
 | 
|  |   1432 | (*This is too strong for version V but would hold for version IV if only B 
 | 
|  |   1433 |   in K6 said a fresh timestamp.
 | 
|  |   1434 | lemma honest_never_says_newer_timestamp:
 | 
|  |   1435 |      "\<lbrakk> (CT evs) \<le> T ; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
 | 
|  |   1436 |      \<Longrightarrow> \<forall> A B. A \<noteq> Spy \<longrightarrow> Says A B X \<notin> set evs"
 | 
|  |   1437 | apply (erule rev_mp)
 | 
|  |   1438 | apply (erule kerbV.induct)
 | 
|  |   1439 | apply (simp_all)
 | 
|  |   1440 | apply force
 | 
|  |   1441 | apply force
 | 
|  |   1442 | txt{*clarifying case K3*}
 | 
|  |   1443 | apply (rule impI)
 | 
|  |   1444 | apply (rule impI)
 | 
|  |   1445 | apply (frule Suc_leD)
 | 
|  |   1446 | apply (clarify)
 | 
|  |   1447 | txt{*cannot solve K3 or K5 because the spy might send CT evs as authTicket
 | 
|  |   1448 | or servTicket, which the honest agent would forward*}
 | 
|  |   1449 | prefer 2 apply force
 | 
|  |   1450 | prefer 4 apply force
 | 
|  |   1451 | prefer 4 apply force
 | 
|  |   1452 | txt{*cannot solve K6 unless B updates the timestamp - rather than bouncing T3*}
 | 
|  |   1453 | oops
 | 
|  |   1454 | *)
 | 
|  |   1455 | 
 | 
|  |   1456 | 
 | 
|  |   1457 | text{*But can prove a less general fact conerning only authenticators!*}
 | 
|  |   1458 | lemma honest_never_says_newer_timestamp_in_auth:
 | 
|  |   1459 |      "\<lbrakk> (CT evs) \<le> T; Number T \<in> parts {X}; A \<notin> bad; evs \<in> kerbV \<rbrakk> 
 | 
|  |   1460 |      \<Longrightarrow> Says A B \<lbrace>Y, X\<rbrace> \<notin> set evs"
 | 
|  |   1461 | apply (erule rev_mp)
 | 
|  |   1462 | apply (erule kerbV.induct)
 | 
|  |   1463 | apply (simp_all)
 | 
|  |   1464 | apply force+
 | 
|  |   1465 | done
 | 
|  |   1466 | 
 | 
|  |   1467 | lemma honest_never_says_current_timestamp_in_auth:
 | 
|  |   1468 |      "\<lbrakk> (CT evs) = T; Number T \<in> parts {X}; A \<notin> bad; evs \<in> kerbV \<rbrakk> 
 | 
|  |   1469 |      \<Longrightarrow> Says A B \<lbrace>Y, X\<rbrace> \<notin> set evs"
 | 
|  |   1470 | apply (frule eq_imp_le)
 | 
|  |   1471 | apply (blast dest: honest_never_says_newer_timestamp_in_auth)
 | 
|  |   1472 | done
 | 
|  |   1473 | 
 | 
|  |   1474 | 
 | 
|  |   1475 | 
 | 
|  |   1476 | lemma A_Issues_B:
 | 
|  |   1477 |      "\<lbrakk> Says A B \<lbrace>ST, Crypt servK \<lbrace>Agent A, Number T3\<rbrace>\<rbrace> \<in> set evs;
 | 
|  |   1478 |          Key servK \<notin> analz (spies evs);
 | 
|  |   1479 |          B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |   1480 |    \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
 | 
|  |   1481 | apply (simp (no_asm) add: Issues_def)
 | 
|  |   1482 | apply (rule exI)
 | 
|  |   1483 | apply (rule conjI, assumption)
 | 
|  |   1484 | apply (simp (no_asm))
 | 
|  |   1485 | apply (erule rev_mp)
 | 
|  |   1486 | apply (erule rev_mp)
 | 
|  |   1487 | apply (erule kerbV.induct, analz_mono_contra)
 | 
|  |   1488 | apply (frule_tac [7] Says_ticket_parts)
 | 
|  |   1489 | apply (frule_tac [5] Says_ticket_parts)
 | 
|  |   1490 | apply (simp_all (no_asm_simp))
 | 
|  |   1491 | txt{*K5*}
 | 
|  |   1492 | apply auto
 | 
|  |   1493 | apply (simp add: takeWhile_tail)
 | 
|  |   1494 | txt{*Level 15: case study necessary because the assumption doesn't state
 | 
|  |   1495 |   the form of servTicket. The guarantee becomes stronger.*}
 | 
|  |   1496 | prefer 2 apply (simp add: takeWhile_tail)
 | 
|  |   1497 | (**This single command of version IV...
 | 
|  |   1498 | apply (blast dest: Says_imp_spies [THEN analz.Inj, THEN analz_Decrypt']
 | 
|  |   1499 |                    K3_imp_K2 K4_trustworthy'
 | 
|  |   1500 |                    parts_spies_takeWhile_mono [THEN subsetD]
 | 
|  |   1501 |                    parts_spies_evs_revD2 [THEN subsetD]
 | 
|  |   1502 |              intro: Says_Auth)
 | 
|  |   1503 | ...expands as follows - including extra exE because of new form of lemmas*)
 | 
|  |   1504 | apply (frule K3_imp_K2, assumption, assumption, erule exE, erule exE)
 | 
|  |   1505 | apply (case_tac "Key authK \<in> analz (spies evs5)")
 | 
|  |   1506 | apply (drule Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Fst, THEN analz_Decrypt', THEN analz.Fst], assumption, assumption, simp)
 | 
|  |   1507 | apply (frule K3_imp_K2, assumption, assumption, erule exE, erule exE)
 | 
|  |   1508 | apply (drule Says_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst])
 | 
|  |   1509 | apply (frule servK_authentic_ter, blast, assumption+)
 | 
|  |   1510 | apply (drule parts_spies_takeWhile_mono [THEN subsetD])
 | 
|  |   1511 | apply (drule parts_spies_evs_revD2 [THEN subsetD])
 | 
|  |   1512 | txt{* @{term Says_K5} closes the proof in version IV because it is clear which 
 | 
|  |   1513 | servTicket an authenticator appears with in msg 5. In version V an authenticator can appear with any item that the spy could replace the servTicket with*}
 | 
|  |   1514 | apply (frule Says_K5, blast, assumption, assumption, assumption, assumption, erule exE)
 | 
|  |   1515 | txt{*We need to state that an honest agent wouldn't send the wrong timestamp
 | 
|  |   1516 | within an authenticator, wathever it is paired with*}
 | 
|  |   1517 | apply (simp add: honest_never_says_current_timestamp_in_auth)
 | 
|  |   1518 | done
 | 
|  |   1519 | 
 | 
|  |   1520 | lemma B_authenticates_and_keydist_to_A:
 | 
|  |   1521 |      "\<lbrakk> Crypt servK \<lbrace>Agent A, Number T3\<rbrace> \<in> parts (spies evs);
 | 
|  |   1522 |          Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>
 | 
|  |   1523 |            \<in> parts (spies evs);
 | 
|  |   1524 |          Key servK \<notin> analz (spies evs);
 | 
|  |   1525 |          B \<noteq> Tgs; A \<notin> bad;  B \<notin> bad;  evs \<in> kerbV \<rbrakk>
 | 
|  |   1526 |    \<Longrightarrow> A Issues B with (Crypt servK \<lbrace>Agent A, Number T3\<rbrace>) on evs"
 | 
|  |   1527 | by (blast dest: B_authenticates_A A_Issues_B)
 | 
|  |   1528 | 
 | 
|  |   1529 | 
 | 
|  |   1530 | 
 | 
|  |   1531 | subsection{*
 | 
|  |   1532 | Novel guarantees, never studied before. Because honest agents always say
 | 
|  |   1533 | the right timestamp in authenticators, we can prove unicity guarantees based 
 | 
|  |   1534 | exactly on timestamps. Classical unicity guarantees are based on nonces.
 | 
|  |   1535 | Of course assuming the agent to be different from the Spy, rather than not in 
 | 
|  |   1536 | bad, would suffice below. Similar guarantees must also hold of
 | 
|  |   1537 | Kerberos IV.*}
 | 
|  |   1538 | 
 | 
|  |   1539 | text{*Notice that an honest agent can send the same timestamp on two
 | 
|  |   1540 | different traces of the same length, but not on the same trace!*}
 | 
|  |   1541 | 
 | 
|  |   1542 | lemma unique_timestamp_authenticator1:
 | 
|  |   1543 |      "\<lbrakk> Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs;
 | 
|  |   1544 |          Says A Kas' \<lbrace>Agent A, Agent Tgs', Number T1\<rbrace> \<in> set evs;
 | 
|  |   1545 |          A \<notin>bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1546 |   \<Longrightarrow> Kas=Kas' \<and> Tgs=Tgs'"
 | 
|  |   1547 | apply (erule rev_mp, erule rev_mp)
 | 
|  |   1548 | apply (erule kerbV.induct)
 | 
|  |   1549 | apply (simp_all, blast)
 | 
|  |   1550 | apply auto
 | 
|  |   1551 | apply (simp_all add: honest_never_says_current_timestamp_in_auth)
 | 
|  |   1552 | done
 | 
|  |   1553 | 
 | 
|  |   1554 | lemma unique_timestamp_authenticator2:
 | 
|  |   1555 |      "\<lbrakk> Says A Tgs \<lbrace>AT, Crypt AK \<lbrace>Agent A, Number T2\<rbrace>, Agent B\<rbrace> \<in> set evs;
 | 
|  |   1556 |      Says A Tgs' \<lbrace>AT', Crypt AK' \<lbrace>Agent A, Number T2\<rbrace>, Agent B'\<rbrace> \<in> set evs;
 | 
|  |   1557 |          A \<notin>bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1558 |   \<Longrightarrow> Tgs=Tgs' \<and> AT=AT' \<and> AK=AK' \<and> B=B'"
 | 
|  |   1559 | apply (erule rev_mp, erule rev_mp)
 | 
|  |   1560 | apply (erule kerbV.induct)
 | 
|  |   1561 | apply (simp_all, blast)
 | 
|  |   1562 | apply auto
 | 
|  |   1563 | apply (simp_all add: honest_never_says_current_timestamp_in_auth)
 | 
|  |   1564 | done
 | 
|  |   1565 | 
 | 
|  |   1566 | lemma unique_timestamp_authenticator3:
 | 
|  |   1567 |      "\<lbrakk> Says A B \<lbrace>ST, Crypt SK \<lbrace>Agent A, Number T\<rbrace>\<rbrace> \<in> set evs;
 | 
|  |   1568 |          Says A B' \<lbrace>ST', Crypt SK' \<lbrace>Agent A, Number T\<rbrace>\<rbrace> \<in> set evs;
 | 
|  |   1569 |          A \<notin>bad; evs \<in> kerbV \<rbrakk>
 | 
|  |   1570 |   \<Longrightarrow> B=B' \<and> ST=ST' \<and> SK=SK'"
 | 
|  |   1571 | apply (erule rev_mp, erule rev_mp)
 | 
|  |   1572 | apply (erule kerbV.induct)
 | 
|  |   1573 | apply (simp_all, blast)
 | 
|  |   1574 | apply (auto simp add: honest_never_says_current_timestamp_in_auth)
 | 
|  |   1575 | done
 | 
|  |   1576 | 
 | 
|  |   1577 | text{*The second part of the message is treated as an authenticator by the last
 | 
|  |   1578 | simplification step, even if it is not an authenticator!*}
 | 
|  |   1579 | lemma unique_timestamp_authticket:
 | 
|  |   1580 |      "\<lbrakk> Says Kas A \<lbrace>X, Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key AK, T\<rbrace>\<rbrace> \<in> set evs;
 | 
|  |   1581 |        Says Kas A' \<lbrace>X', Crypt (shrK Tgs') \<lbrace>Agent A', Agent Tgs', Key AK', T\<rbrace>\<rbrace> \<in> set evs;
 | 
|  |   1582 |          evs \<in> kerbV \<rbrakk>
 | 
|  |   1583 |   \<Longrightarrow> A=A' \<and> X=X' \<and> Tgs=Tgs' \<and> AK=AK'"
 | 
|  |   1584 | apply (erule rev_mp, erule rev_mp)
 | 
|  |   1585 | apply (erule kerbV.induct)
 | 
|  |   1586 | apply (auto simp add: honest_never_says_current_timestamp_in_auth)
 | 
|  |   1587 | done
 | 
|  |   1588 | 
 | 
|  |   1589 | text{*The second part of the message is treated as an authenticator by the last
 | 
|  |   1590 | simplification step, even if it is not an authenticator!*}
 | 
|  |   1591 | lemma unique_timestamp_servticket:
 | 
|  |   1592 |      "\<lbrakk> Says Tgs A \<lbrace>X, Crypt (shrK B) \<lbrace>Agent A, Agent B, Key SK, T\<rbrace>\<rbrace> \<in> set evs;
 | 
|  |   1593 |        Says Tgs A' \<lbrace>X', Crypt (shrK B') \<lbrace>Agent A', Agent B', Key SK', T\<rbrace>\<rbrace> \<in> set evs;
 | 
|  |   1594 |          evs \<in> kerbV \<rbrakk>
 | 
|  |   1595 |   \<Longrightarrow> A=A' \<and> X=X' \<and> B=B' \<and> SK=SK'"
 | 
|  |   1596 | apply (erule rev_mp, erule rev_mp)
 | 
|  |   1597 | apply (erule kerbV.induct)
 | 
|  |   1598 | apply (auto simp add: honest_never_says_current_timestamp_in_auth)
 | 
|  |   1599 | done
 | 
|  |   1600 | 
 | 
|  |   1601 | (*Uses assumption K6's assumption that B \<noteq> Kas, otherwise B should say
 | 
|  |   1602 | fresh timestamp*)
 | 
|  |   1603 | lemma Kas_never_says_newer_timestamp:
 | 
|  |   1604 |      "\<lbrakk> (CT evs) \<le> T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
 | 
|  |   1605 |      \<Longrightarrow> \<forall> A. Says Kas A X \<notin> set evs"
 | 
|  |   1606 | apply (erule rev_mp)
 | 
|  |   1607 | apply (erule kerbV.induct, auto)
 | 
|  |   1608 | done
 | 
|  |   1609 | 
 | 
|  |   1610 | lemma Kas_never_says_current_timestamp:
 | 
|  |   1611 |      "\<lbrakk> (CT evs) = T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
 | 
|  |   1612 |      \<Longrightarrow> \<forall> A. Says Kas A X \<notin> set evs"
 | 
|  |   1613 | apply (frule eq_imp_le)
 | 
|  |   1614 | apply (blast dest: Kas_never_says_newer_timestamp)
 | 
|  |   1615 | done
 | 
|  |   1616 | 
 | 
|  |   1617 | lemma unique_timestamp_msg2:
 | 
|  |   1618 |      "\<lbrakk> Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key AK, Agent Tgs, T\<rbrace>, AT\<rbrace> \<in> set evs;
 | 
|  |   1619 |      Says Kas A' \<lbrace>Crypt (shrK A') \<lbrace>Key AK', Agent Tgs', T\<rbrace>, AT'\<rbrace> \<in> set evs;
 | 
|  |   1620 |          evs \<in> kerbV \<rbrakk>
 | 
|  |   1621 |   \<Longrightarrow> A=A' \<and> AK=AK' \<and> Tgs=Tgs' \<and> AT=AT'"
 | 
|  |   1622 | apply (erule rev_mp, erule rev_mp)
 | 
|  |   1623 | apply (erule kerbV.induct)
 | 
|  |   1624 | apply (auto simp add: Kas_never_says_current_timestamp)
 | 
|  |   1625 | done
 | 
|  |   1626 | 
 | 
|  |   1627 | (*Uses assumption K6's assumption that B \<noteq> Tgs, otherwise B should say
 | 
|  |   1628 | fresh timestamp*)
 | 
|  |   1629 | lemma Tgs_never_says_newer_timestamp:
 | 
|  |   1630 |      "\<lbrakk> (CT evs) \<le> T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
 | 
|  |   1631 |      \<Longrightarrow> \<forall> A. Says Tgs A X \<notin> set evs"
 | 
|  |   1632 | apply (erule rev_mp)
 | 
|  |   1633 | apply (erule kerbV.induct, auto)
 | 
|  |   1634 | done
 | 
|  |   1635 | 
 | 
|  |   1636 | lemma Tgs_never_says_current_timestamp:
 | 
|  |   1637 |      "\<lbrakk> (CT evs) = T; Number T \<in> parts {X}; evs \<in> kerbV \<rbrakk> 
 | 
|  |   1638 |      \<Longrightarrow> \<forall> A. Says Tgs A X \<notin> set evs"
 | 
|  |   1639 | apply (frule eq_imp_le)
 | 
|  |   1640 | apply (blast dest: Tgs_never_says_newer_timestamp)
 | 
|  |   1641 | done
 | 
|  |   1642 | 
 | 
|  |   1643 | 
 | 
|  |   1644 | lemma unique_timestamp_msg4:
 | 
|  |   1645 |      "\<lbrakk> Says Tgs A \<lbrace>Crypt (shrK A) \<lbrace>Key SK, Agent B, T\<rbrace>, ST\<rbrace> \<in> set evs;
 | 
|  |   1646 |        Says Tgs A' \<lbrace>Crypt (shrK A') \<lbrace>Key SK', Agent B', T\<rbrace>, ST'\<rbrace> \<in> set evs;
 | 
|  |   1647 |          evs \<in> kerbV \<rbrakk> 
 | 
|  |   1648 |   \<Longrightarrow> A=A' \<and> SK=SK' \<and> B=B' \<and> ST=ST'"
 | 
|  |   1649 | apply (erule rev_mp, erule rev_mp)
 | 
|  |   1650 | apply (erule kerbV.induct)
 | 
|  |   1651 | apply (auto simp add: Tgs_never_says_current_timestamp)
 | 
|  |   1652 | done
 | 
|  |   1653 |  
 | 
|  |   1654 | end
 |