src/HOL/Auth/KerberosIV.ML
changeset 6452 6a1b393ccdc0
child 7239 26685edee372
equal deleted inserted replaced
6451:bc943acc5fda 6452:6a1b393ccdc0
       
     1 (*  Title:      HOL/Auth/KerberosIV
       
     2     ID:         $Id$
       
     3     Author:     Giampaolo Bella, Cambridge University Computer Laboratory
       
     4     Copyright   1998  University of Cambridge
       
     5 
       
     6 The Kerberos protocol, version IV.
       
     7 *)
       
     8 
       
     9 Pretty.setdepth 20;
       
    10 proof_timing:=true;
       
    11 HOL_quantifiers := false;
       
    12 
       
    13 AddIffs [AuthLife_LB, ServLife_LB, AutcLife_LB, RespLife_LB, Tgs_not_bad];
       
    14 
       
    15 
       
    16 (** Reversed traces **)
       
    17 
       
    18 Goal "spies (evs @ [Says A B X]) = insert X (spies evs)";
       
    19 by (induct_tac "evs" 1);
       
    20 by (induct_tac "a" 2);
       
    21 by Auto_tac;
       
    22 qed "spies_Says_rev";
       
    23 
       
    24 Goal "spies (evs @ [Gets A X]) = spies evs";
       
    25 by (induct_tac "evs" 1);
       
    26 by (induct_tac "a" 2);
       
    27 by Auto_tac;
       
    28 qed "spies_Gets_rev";
       
    29 
       
    30 Goal "spies (evs @ [Notes A X]) = \
       
    31 \         (if A:bad then insert X (spies evs) else spies evs)";
       
    32 by (induct_tac "evs" 1);
       
    33 by (induct_tac "a" 2);
       
    34 by Auto_tac;
       
    35 qed "spies_Notes_rev";
       
    36 
       
    37 Goal "spies evs = spies (rev evs)";
       
    38 by (induct_tac "evs" 1);
       
    39 by (induct_tac "a" 2);
       
    40 by (ALLGOALS 
       
    41     (asm_simp_tac (simpset() addsimps [spies_Says_rev, spies_Gets_rev, 
       
    42 				       spies_Notes_rev])));
       
    43 qed "spies_evs_rev";
       
    44 bind_thm ("parts_spies_evs_revD2", spies_evs_rev RS equalityD2 RS parts_mono);
       
    45 
       
    46 Goal "spies (takeWhile P evs)  <=  spies evs";
       
    47 by (induct_tac "evs" 1);
       
    48 by (induct_tac "a" 2);
       
    49 by Auto_tac;
       
    50 (* Resembles "used_subset_append" in Event.ML*)
       
    51 qed "spies_takeWhile";
       
    52 bind_thm ("parts_spies_takeWhile_mono", spies_takeWhile RS parts_mono);
       
    53 
       
    54 Goal "~P(x) --> takeWhile P (xs @ [x]) = takeWhile P xs";
       
    55 by (induct_tac "xs" 1);
       
    56 by Auto_tac;
       
    57 qed "takeWhile_tail";
       
    58 
       
    59 
       
    60 (*****************LEMMAS ABOUT AuthKeys****************)
       
    61 
       
    62 Goalw [AuthKeys_def] "AuthKeys [] = {}";
       
    63 by (Simp_tac 1);
       
    64 qed "AuthKeys_empty";
       
    65 
       
    66 Goalw [AuthKeys_def] 
       
    67  "(ALL A Tk akey Peer.              \
       
    68 \  ev ~= Says Kas A (Crypt (shrK A) {|akey, Agent Peer, Tk,      \
       
    69 \             (Crypt (shrK Peer) {|Agent A, Agent Peer, akey, Tk|})|}))\ 
       
    70 \      ==> AuthKeys (ev # evs) = AuthKeys evs";
       
    71 by Auto_tac;
       
    72 qed "AuthKeys_not_insert";
       
    73 
       
    74 Goalw [AuthKeys_def] 
       
    75   "AuthKeys \
       
    76 \    (Says Kas A (Crypt (shrK A) {|Key K, Agent Peer, Number Tk, \
       
    77 \     (Crypt (shrK Peer) {|Agent A, Agent Peer, Key K, Number Tk|})|}) # evs) \
       
    78 \      = insert K (AuthKeys evs)";
       
    79 by Auto_tac;
       
    80 qed "AuthKeys_insert";
       
    81 
       
    82 Goalw [AuthKeys_def] 
       
    83    "K : AuthKeys \
       
    84 \   (Says Kas A (Crypt (shrK A) {|Key K', Agent Peer, Number Tk, \
       
    85 \    (Crypt (shrK Peer) {|Agent A, Agent Peer, Key K', Number Tk|})|}) # evs) \
       
    86 \       ==> K = K' | K : AuthKeys evs";
       
    87 by Auto_tac;
       
    88 qed "AuthKeys_simp";
       
    89 
       
    90 Goalw [AuthKeys_def]  
       
    91    "Says Kas A (Crypt (shrK A) {|Key K, Agent Tgs, Number Tk, \
       
    92 \    (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key K, Number Tk|})|}) : set evs \
       
    93 \       ==> K : AuthKeys evs";
       
    94 by Auto_tac;
       
    95 qed "AuthKeysI";
       
    96 
       
    97 Goalw [AuthKeys_def] "K : AuthKeys evs ==> Key K : used evs";
       
    98 by (Simp_tac 1);
       
    99 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   100 qed "AuthKeys_used";
       
   101 
       
   102 
       
   103 (**** FORWARDING LEMMAS ****)
       
   104 
       
   105 (*--For reasoning about the encrypted portion of message K3--*)
       
   106 Goal "Says Kas' A (Crypt KeyA {|AuthKey, Peer, Tk, AuthTicket|}) \
       
   107 \              : set evs ==> AuthTicket : parts (spies evs)";
       
   108 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   109 qed "K3_msg_in_parts_spies";
       
   110 
       
   111 Goal "Says Kas A (Crypt KeyA {|AuthKey, Peer, Tk, AuthTicket|}) \
       
   112 \              : set evs ==> AuthKey : parts (spies evs)";
       
   113 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   114 qed "Oops_parts_spies1";
       
   115                               
       
   116 Goal "[| Says Kas A (Crypt KeyA {|Key AuthKey, Peer, Tk, AuthTicket|}) \
       
   117 \          : set evs ;\
       
   118 \        evs : kerberos |] ==> AuthKey ~: range shrK";
       
   119 by (etac rev_mp 1);
       
   120 by (etac kerberos.induct 1);
       
   121 by Auto_tac;
       
   122 qed "Oops_range_spies1";
       
   123 
       
   124 (*--For reasoning about the encrypted portion of message K5--*)
       
   125 Goal "Says Tgs' A (Crypt AuthKey {|ServKey, Agent B, Tt, ServTicket|})\
       
   126  \             : set evs ==> ServTicket : parts (spies evs)";
       
   127 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   128 qed "K5_msg_in_parts_spies";
       
   129 
       
   130 Goal "Says Tgs A (Crypt AuthKey {|ServKey, Agent B, Tt, ServTicket|})\
       
   131 \                  : set evs ==> ServKey : parts (spies evs)";
       
   132 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   133 qed "Oops_parts_spies2";
       
   134 
       
   135 Goal "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}) \
       
   136 \          : set evs ;\
       
   137 \        evs : kerberos |] ==> ServKey ~: range shrK";
       
   138 by (etac rev_mp 1);
       
   139 by (etac kerberos.induct 1);
       
   140 by Auto_tac;
       
   141 qed "Oops_range_spies2";
       
   142 
       
   143 Goal "Says S A (Crypt K {|SesKey, B, TimeStamp, Ticket|}) : set evs \
       
   144 \     ==> Ticket : parts (spies evs)";
       
   145 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   146 qed "Says_ticket_in_parts_spies";
       
   147 (*Replaces both K3_msg_in_parts_spies and K5_msg_in_parts_spies*)
       
   148 
       
   149 fun parts_induct_tac i = 
       
   150     etac kerberos.induct i  THEN 
       
   151     REPEAT (FIRSTGOAL analz_mono_contra_tac)  THEN
       
   152     forward_tac [K3_msg_in_parts_spies] (i+4)  THEN
       
   153     forward_tac [K5_msg_in_parts_spies] (i+6)  THEN
       
   154     forward_tac [Oops_parts_spies1] (i+8)  THEN
       
   155     forward_tac [Oops_parts_spies2] (i+9) THEN
       
   156     prove_simple_subgoals_tac 1;
       
   157 
       
   158 
       
   159 (*Spy never sees another agent's shared key! (unless it's lost at start)*)
       
   160 Goal "evs : kerberos ==> (Key (shrK A) : parts (spies evs)) = (A : bad)";
       
   161 by (parts_induct_tac 1);
       
   162 by (Fake_parts_insert_tac 1);
       
   163 by (ALLGOALS Blast_tac);
       
   164 qed "Spy_see_shrK";
       
   165 Addsimps [Spy_see_shrK];
       
   166 
       
   167 Goal "evs : kerberos ==> (Key (shrK A) : analz (spies evs)) = (A : bad)";
       
   168 by (auto_tac (claset() addDs [impOfSubs analz_subset_parts], simpset()));
       
   169 qed "Spy_analz_shrK";
       
   170 Addsimps [Spy_analz_shrK];
       
   171 
       
   172 Goal "[| Key (shrK A) : parts (spies evs);  evs : kerberos |] ==> A:bad";
       
   173 by (blast_tac (claset() addDs [Spy_see_shrK]) 1);
       
   174 qed "Spy_see_shrK_D";
       
   175 bind_thm ("Spy_analz_shrK_D", analz_subset_parts RS subsetD RS Spy_see_shrK_D);
       
   176 AddSDs [Spy_see_shrK_D, Spy_analz_shrK_D];
       
   177 
       
   178 (*Nobody can have used non-existent keys!*)
       
   179 Goal "evs : kerberos ==>      \
       
   180 \     Key K ~: used evs --> K ~: keysFor (parts (spies evs))";
       
   181 by (parts_induct_tac 1);
       
   182 (*Fake*)
       
   183 by (best_tac
       
   184       (claset() addSDs [impOfSubs (parts_insert_subset_Un RS keysFor_mono)]
       
   185                addIs  [impOfSubs analz_subset_parts]
       
   186                addDs  [impOfSubs (analz_subset_parts RS keysFor_mono)]
       
   187                addss  (simpset())) 1);
       
   188 (*Others*)
       
   189 by (ALLGOALS (blast_tac (claset() addSEs spies_partsEs)));
       
   190 qed_spec_mp "new_keys_not_used";
       
   191 
       
   192 bind_thm ("new_keys_not_analzd",
       
   193           [analz_subset_parts RS keysFor_mono,
       
   194            new_keys_not_used] MRS contra_subsetD);
       
   195 
       
   196 Addsimps [new_keys_not_used, new_keys_not_analzd];
       
   197 
       
   198 
       
   199 (*********************** REGULARITY LEMMAS ***********************)
       
   200 (*       concerning the form of items passed in messages         *)
       
   201 (*****************************************************************)
       
   202 
       
   203 (*Describes the form of AuthKey, AuthTicket, and K sent by Kas*)
       
   204 Goal "[| Says Kas A (Crypt K {|Key AuthKey, Agent Peer, Tk, AuthTicket|}) \
       
   205 \          : set evs;                 \
       
   206 \        evs : kerberos |]             \
       
   207 \     ==> AuthKey ~: range shrK & AuthKey : AuthKeys evs & \
       
   208 \ AuthTicket = (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|} ) &\
       
   209 \            K = shrK A  & Peer = Tgs";
       
   210 by (etac rev_mp 1);
       
   211 by (etac kerberos.induct 1);
       
   212 by (ALLGOALS (simp_tac (simpset() addsimps [AuthKeys_def, AuthKeys_insert])));
       
   213 by (ALLGOALS Blast_tac);
       
   214 qed "Says_Kas_message_form";
       
   215 
       
   216 (*This lemma is essential for proving Says_Tgs_message_form: 
       
   217   
       
   218   the session key AuthKey
       
   219   supplied by Kas in the authentication ticket
       
   220   cannot be a long-term key!
       
   221 
       
   222   Generalised to any session keys (both AuthKey and ServKey).
       
   223 *)
       
   224 Goal "[| Crypt (shrK Tgs_B) {|Agent A, Agent Tgs_B, Key SesKey, Number T|}\
       
   225 \           : parts (spies evs); Tgs_B ~: bad;\
       
   226 \        evs : kerberos |]    \
       
   227 \     ==> SesKey ~: range shrK";
       
   228 by (etac rev_mp 1);
       
   229 by (parts_induct_tac 1);
       
   230 by (Fake_parts_insert_tac 1);
       
   231 qed "SesKey_is_session_key";
       
   232 
       
   233 Goal "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}  \
       
   234 \          : parts (spies evs);                              \
       
   235 \        evs : kerberos |]                          \
       
   236 \     ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, \
       
   237 \                Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}|})  \
       
   238 \           : set evs";
       
   239 by (etac rev_mp 1);
       
   240 by (parts_induct_tac 1);
       
   241 (*Fake*)
       
   242 by (Fake_parts_insert_tac 1);
       
   243 (*K4*)
       
   244 by (Blast_tac 1);
       
   245 qed "A_trusts_AuthTicket";
       
   246 
       
   247 Goal "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}\
       
   248 \          : parts (spies evs);\
       
   249 \        evs : kerberos |]    \
       
   250 \     ==> AuthKey : AuthKeys evs";
       
   251 by (forward_tac [A_trusts_AuthTicket] 1);
       
   252 by (assume_tac 1);
       
   253 by (simp_tac (simpset() addsimps [AuthKeys_def]) 1);
       
   254 by (Blast_tac 1);
       
   255 qed "AuthTicket_crypt_AuthKey";
       
   256 
       
   257 (*Describes the form of ServKey, ServTicket and AuthKey sent by Tgs*)
       
   258 Goal "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})\
       
   259 \          : set evs; \
       
   260 \        evs : kerberos |]    \
       
   261 \  ==> B ~= Tgs & ServKey ~: range shrK & ServKey ~: AuthKeys evs &\
       
   262 \      ServTicket = (Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|} ) & \
       
   263 \      AuthKey ~: range shrK & AuthKey : AuthKeys evs";
       
   264 by (etac rev_mp 1);
       
   265 by (etac kerberos.induct 1);
       
   266 by (ALLGOALS
       
   267     (asm_full_simp_tac
       
   268      (simpset() addsimps [AuthKeys_insert, AuthKeys_not_insert,
       
   269 			  AuthKeys_empty, AuthKeys_simp])));
       
   270 by (blast_tac (claset() addEs  spies_partsEs) 1);
       
   271 by Auto_tac;
       
   272 by (blast_tac (claset() addSDs [AuthKeys_used, Says_Kas_message_form]) 1);
       
   273 by (blast_tac (claset() addSDs [SesKey_is_session_key]
       
   274                         addEs spies_partsEs) 1);
       
   275 by (blast_tac (claset() addDs [AuthTicket_crypt_AuthKey] 
       
   276                         addEs spies_partsEs) 1);
       
   277 qed "Says_Tgs_message_form";
       
   278 
       
   279 (*If a certain encrypted message appears then it originated with Kas*)
       
   280 Goal "[| Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|}  \
       
   281 \          : parts (spies evs);                              \
       
   282 \        A ~: bad;  evs : kerberos |]                        \
       
   283 \     ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|})  \
       
   284 \           : set evs";
       
   285 by (etac rev_mp 1);
       
   286 by (parts_induct_tac 1);
       
   287 (*Fake*)
       
   288 by (Fake_parts_insert_tac 1);
       
   289 (*K4*)
       
   290 by (blast_tac (claset() addSDs [Says_imp_spies RS parts.Inj RS parts.Fst,
       
   291 			       A_trusts_AuthTicket RS Says_Kas_message_form])
       
   292     1);
       
   293 qed "A_trusts_AuthKey";
       
   294 
       
   295 
       
   296 (*If a certain encrypted message appears then it originated with Tgs*)
       
   297 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}     \
       
   298 \          : parts (spies evs);                                     \
       
   299 \        Key AuthKey ~: analz (spies evs);           \
       
   300 \        AuthKey ~: range shrK;                      \
       
   301 \        evs : kerberos |]         \
       
   302 \==> EX A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})\
       
   303 \      : set evs";
       
   304 by (etac rev_mp 1);
       
   305 by (etac rev_mp 1);
       
   306 by (parts_induct_tac 1);
       
   307 (*Fake*)
       
   308 by (Fake_parts_insert_tac 1);
       
   309 (*K2*)
       
   310 by (Blast_tac 1);
       
   311 (*K4*)
       
   312 by Auto_tac;
       
   313 qed "A_trusts_K4";
       
   314 
       
   315 Goal "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|} \
       
   316 \          : parts (spies evs);          \
       
   317 \        A ~: bad;                       \
       
   318 \        evs : kerberos |]                \
       
   319 \   ==> AuthKey ~: range shrK &               \
       
   320 \       AuthTicket = Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}";
       
   321 by (etac rev_mp 1);
       
   322 by (parts_induct_tac 1);
       
   323 by (Fake_parts_insert_tac 1);
       
   324 by (Blast_tac 1);
       
   325 qed "AuthTicket_form";
       
   326 
       
   327 (* This form holds also over an AuthTicket, but is not needed below.     *)
       
   328 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|} \
       
   329 \             : parts (spies evs); \
       
   330 \           Key AuthKey ~: analz (spies evs);  \
       
   331 \           evs : kerberos |]                                       \
       
   332 \        ==> ServKey ~: range shrK &  \
       
   333 \   (EX A. ServTicket = Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})";
       
   334 by (etac rev_mp 1);
       
   335 by (etac rev_mp 1);
       
   336 by (parts_induct_tac 1);
       
   337 by (Fake_parts_insert_tac 1);
       
   338 qed "ServTicket_form";
       
   339 
       
   340 Goal "[| Says Kas' A (Crypt (shrK A) \
       
   341 \             {|Key AuthKey, Agent Tgs, Tk, AuthTicket|} ) : set evs; \
       
   342 \        evs : kerberos |]    \
       
   343 \     ==> AuthKey ~: range shrK & \
       
   344 \         AuthTicket = \
       
   345 \                 Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}\
       
   346 \         | AuthTicket : analz (spies evs)";
       
   347 by (case_tac "A : bad" 1);
       
   348 by (force_tac (claset() addSDs [Says_imp_spies RS analz.Inj], simpset()) 1);
       
   349 by (forward_tac [Says_imp_spies RS parts.Inj] 1);
       
   350 by (blast_tac (claset() addSDs [AuthTicket_form]) 1);
       
   351 qed "Says_kas_message_form";
       
   352 (* Essentially the same as AuthTicket_form *)
       
   353 
       
   354 Goal "[| Says Tgs' A (Crypt AuthKey \
       
   355 \             {|Key ServKey, Agent B, Tt, ServTicket|} ) : set evs; \
       
   356 \        evs : kerberos |]    \
       
   357 \     ==> ServKey ~: range shrK & \
       
   358 \         (EX A. ServTicket = \
       
   359 \                 Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})  \
       
   360 \          | ServTicket : analz (spies evs)";
       
   361 by (case_tac "Key AuthKey : analz (spies evs)" 1);
       
   362 by (force_tac (claset() addSDs [Says_imp_spies RS analz.Inj], simpset()) 1);
       
   363 by (forward_tac [Says_imp_spies RS parts.Inj] 1);
       
   364 by (blast_tac (claset() addSDs [ServTicket_form]) 1);
       
   365 qed "Says_tgs_message_form";
       
   366 (* This form MUST be used in analz_sees_tac below *)
       
   367 
       
   368 
       
   369 (*****************UNICITY THEOREMS****************)
       
   370 
       
   371 (* The session key, if secure, uniquely identifies the Ticket
       
   372    whether AuthTicket or ServTicket. As a matter of fact, one can read
       
   373    also Tgs in the place of B.                                     *)
       
   374 Goal "evs : kerberos ==>                                        \
       
   375 \     Key SesKey ~: analz (spies evs) -->   \
       
   376 \     (EX A B T. ALL A' B' T'.                          \
       
   377 \      Crypt (shrK B') {|Agent A', Agent B', Key SesKey, T'|}    \
       
   378 \        : parts (spies evs) --> A'=A & B'=B & T'=T)";
       
   379 by (parts_induct_tac 1);
       
   380 by (REPEAT (etac (exI RSN (2,exE)) 1)   (*stripping EXs makes proof faster*)
       
   381     THEN Fake_parts_insert_tac 1);
       
   382 by (asm_simp_tac (simpset() addsimps [all_conj_distrib]) 1); 
       
   383 by (expand_case_tac "SesKey = ?y" 1);
       
   384 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   385 by (expand_case_tac "SesKey = ?y" 1);
       
   386 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   387 val lemma = result();
       
   388 
       
   389 Goal "[| Crypt (shrK B)  {|Agent A,  Agent B,  Key SesKey, T|}        \
       
   390 \          : parts (spies evs);            \
       
   391 \        Crypt (shrK B') {|Agent A', Agent B', Key SesKey, T'|}     \
       
   392 \          : parts (spies evs);            \
       
   393 \        evs : kerberos;  Key SesKey ~: analz (spies evs) |]  \
       
   394 \     ==> A=A' & B=B' & T=T'";
       
   395 by (prove_unique_tac lemma 1);
       
   396 qed "unique_CryptKey";
       
   397 
       
   398 Goal "evs : kerberos \
       
   399 \     ==> Key SesKey ~: analz (spies evs) -->   \
       
   400 \         (EX K' B' T' Ticket'. ALL K B T Ticket.                          \
       
   401 \          Crypt K {|Key SesKey, Agent B, T, Ticket|}    \
       
   402 \            : parts (spies evs) --> K=K' & B=B' & T=T' & Ticket=Ticket')";
       
   403 by (parts_induct_tac 1);
       
   404 by (REPEAT (etac (exI RSN (2,exE)) 1) THEN Fake_parts_insert_tac 1);
       
   405 by (asm_simp_tac (simpset() addsimps [all_conj_distrib]) 1); 
       
   406 by (expand_case_tac "SesKey = ?y" 1);
       
   407 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   408 by (expand_case_tac "SesKey = ?y" 1);
       
   409 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   410 val lemma = result();
       
   411 
       
   412 (*An AuthKey is encrypted by one and only one Shared key.
       
   413   A ServKey is encrypted by one and only one AuthKey.
       
   414 *)
       
   415 Goal "[| Crypt K  {|Key SesKey,  Agent B, T, Ticket|}        \
       
   416 \          : parts (spies evs);            \
       
   417 \        Crypt K' {|Key SesKey,  Agent B', T', Ticket'|}     \
       
   418 \          : parts (spies evs);            \
       
   419 \        evs : kerberos;  Key SesKey ~: analz (spies evs) |]  \
       
   420 \     ==> K=K' & B=B' & T=T' & Ticket=Ticket'";
       
   421 by (prove_unique_tac lemma 1);
       
   422 qed "Key_unique_SesKey";
       
   423 
       
   424 
       
   425 (*
       
   426   At reception of any message mentioning A, Kas associates shrK A with
       
   427   a new AuthKey. Realistic, as the user gets a new AuthKey at each login.
       
   428   Similarly, at reception of any message mentioning an AuthKey
       
   429   (a legitimate user could make several requests to Tgs - by K3), Tgs 
       
   430   associates it with a new ServKey.
       
   431 
       
   432   Therefore, a goal like
       
   433 
       
   434    "evs : kerberos \
       
   435   \  ==> Key Kc ~: analz (spies evs) -->   \
       
   436   \        (EX K' B' T' Ticket'. ALL K B T Ticket.                          \
       
   437   \         Crypt Kc {|Key K, Agent B, T, Ticket|}    \
       
   438   \          : parts (spies evs) --> K=K' & B=B' & T=T' & Ticket=Ticket')";
       
   439 
       
   440   would fail on the K2 and K4 cases.
       
   441 *)
       
   442 
       
   443 (* AuthKey uniquely identifies the message from Kas *)
       
   444 Goal "evs : kerberos ==>                                        \
       
   445 \      EX A' Ka' Tk' X'. ALL A Ka Tk X.                          \
       
   446 \        Says Kas A (Crypt Ka {|Key AuthKey, Agent Tgs, Tk, X|})  \
       
   447 \          : set evs --> A=A' & Ka=Ka' & Tk=Tk' & X=X'";
       
   448 by (etac kerberos.induct 1);
       
   449 by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib])));
       
   450 by (Step_tac 1);
       
   451 (*K2: it can't be a new key*)
       
   452 by (expand_case_tac "AuthKey = ?y" 1);
       
   453 by (REPEAT (ares_tac [refl,exI,impI,conjI] 2));
       
   454 by (blast_tac (claset() delrules [conjI]   (*prevent split-up into 4 subgoals*)
       
   455                       addSEs spies_partsEs) 1);
       
   456 val lemma = result();
       
   457 
       
   458 Goal "[| Says Kas A                                          \
       
   459 \             (Crypt Ka {|Key AuthKey, Agent Tgs, Tk, X|}) : set evs;     \ 
       
   460 \        Says Kas A'                                         \
       
   461 \             (Crypt Ka' {|Key AuthKey, Agent Tgs, Tk', X'|}) : set evs;   \
       
   462 \        evs : kerberos |] ==> A=A' & Ka=Ka' & Tk=Tk' & X=X'";
       
   463 by (prove_unique_tac lemma 1);
       
   464 qed "unique_AuthKeys";
       
   465 
       
   466 (* ServKey uniquely identifies the message from Tgs *)
       
   467 Goal "evs : kerberos ==>                                        \
       
   468 \      EX A' B' AuthKey' Tk' X'. ALL A B AuthKey Tk X.                     \
       
   469 \        Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tk, X|})  \
       
   470 \          : set evs --> A=A' & B=B' & AuthKey=AuthKey' & Tk=Tk' & X=X'";
       
   471 by (etac kerberos.induct 1);
       
   472 by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib])));
       
   473 by (Step_tac 1);
       
   474 (*K4: it can't be a new key*)
       
   475 by (expand_case_tac "ServKey = ?y" 1);
       
   476 by (REPEAT (ares_tac [refl,exI,impI,conjI] 2));
       
   477 by (blast_tac (claset() delrules [conjI]   (*prevent split-up into 4 subgoals*)
       
   478                       addSEs spies_partsEs) 1);
       
   479 val lemma = result();
       
   480 
       
   481 Goal "[| Says Tgs A                                             \
       
   482 \             (Crypt K {|Key ServKey, Agent B, Tt, X|}) : set evs; \ 
       
   483 \        Says Tgs A'                                                 \
       
   484 \             (Crypt K' {|Key ServKey, Agent B', Tt', X'|}) : set evs; \
       
   485 \        evs : kerberos |] ==> A=A' & B=B' & K=K' & Tt=Tt' & X=X'";
       
   486 by (prove_unique_tac lemma 1);
       
   487 qed "unique_ServKeys";
       
   488 
       
   489 
       
   490 (*****************LEMMAS ABOUT the predicate KeyCryptKey****************)
       
   491 
       
   492 Goalw [KeyCryptKey_def] "~ KeyCryptKey AuthKey ServKey []";
       
   493 by (Simp_tac 1);
       
   494 qed "not_KeyCryptKey_Nil";
       
   495 AddIffs [not_KeyCryptKey_Nil];
       
   496 
       
   497 Goalw [KeyCryptKey_def]
       
   498  "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |}) \
       
   499 \             : set evs;    \
       
   500 \           evs : kerberos |] ==> KeyCryptKey AuthKey ServKey evs";
       
   501 by (forward_tac [Says_Tgs_message_form] 1);
       
   502 by (assume_tac 1);
       
   503 by (Blast_tac 1);
       
   504 qed "KeyCryptKeyI";
       
   505 
       
   506 Goalw [KeyCryptKey_def]
       
   507    "KeyCryptKey AuthKey ServKey (Says S A X # evs) =                       \
       
   508 \    (Tgs = S &                                                            \
       
   509 \     (EX B tt. X = Crypt AuthKey        \
       
   510 \               {|Key ServKey, Agent B, tt,  \
       
   511 \                 Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} |}) \
       
   512 \    | KeyCryptKey AuthKey ServKey evs)";
       
   513 by (Simp_tac 1);
       
   514 by (Blast_tac 1);
       
   515 qed "KeyCryptKey_Says";
       
   516 Addsimps [KeyCryptKey_Says];
       
   517 
       
   518 (*A fresh AuthKey cannot be associated with any other
       
   519   (with respect to a given trace). *)
       
   520 Goalw [KeyCryptKey_def]
       
   521  "[| Key AuthKey ~: used evs; evs : kerberos |] \
       
   522 \        ==> ~ KeyCryptKey AuthKey ServKey evs";
       
   523 by (etac rev_mp 1);
       
   524 by (parts_induct_tac 1);
       
   525 by (Asm_full_simp_tac 1);
       
   526 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   527 qed "Auth_fresh_not_KeyCryptKey";
       
   528 
       
   529 (*A fresh ServKey cannot be associated with any other
       
   530   (with respect to a given trace). *)
       
   531 Goalw [KeyCryptKey_def]
       
   532  "Key ServKey ~: used evs ==> ~ KeyCryptKey AuthKey ServKey evs";
       
   533 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   534 qed "Serv_fresh_not_KeyCryptKey";
       
   535 
       
   536 Goalw [KeyCryptKey_def]
       
   537  "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, tk|}\
       
   538 \             : parts (spies evs);  evs : kerberos |] \
       
   539 \        ==> ~ KeyCryptKey K AuthKey evs";
       
   540 by (etac rev_mp 1);
       
   541 by (parts_induct_tac 1);
       
   542 (*K4*)
       
   543 by (blast_tac (claset() addEs spies_partsEs) 3);
       
   544 (*K2: by freshness*)
       
   545 by (blast_tac (claset() addEs spies_partsEs) 2);
       
   546 by (Fake_parts_insert_tac 1);
       
   547 qed "AuthKey_not_KeyCryptKey";
       
   548 
       
   549 (*A secure serverkey cannot have been used to encrypt others*)
       
   550 Goalw [KeyCryptKey_def]
       
   551  "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} \
       
   552 \             : parts (spies evs);                     \
       
   553 \           Key ServKey ~: analz (spies evs);             \
       
   554 \           B ~= Tgs;  evs : kerberos |] \
       
   555 \        ==> ~ KeyCryptKey ServKey K evs";
       
   556 by (etac rev_mp 1);
       
   557 by (etac rev_mp 1);
       
   558 by (parts_induct_tac 1);
       
   559 by (Fake_parts_insert_tac 1);
       
   560 (*K4 splits into distinct subcases*)
       
   561 by (Step_tac 1);
       
   562 by (ALLGOALS Asm_full_simp_tac);
       
   563 (*ServKey can't have been enclosed in two certificates*)
       
   564 by (blast_tac (claset() addSDs [Says_imp_spies RS parts.Inj]
       
   565                        addSEs [MPair_parts]
       
   566                        addDs  [unique_CryptKey]) 4);
       
   567 (*ServKey is fresh and so could not have been used, by new_keys_not_used*)
       
   568 by (force_tac (claset() addSDs [Says_imp_spies RS parts.Inj,
       
   569 				Crypt_imp_invKey_keysFor],
       
   570 	       simpset()) 2); 
       
   571 (*Others by freshness*)
       
   572 by (REPEAT (blast_tac (claset() addSEs spies_partsEs) 1));
       
   573 qed "ServKey_not_KeyCryptKey";
       
   574 
       
   575 (*Long term keys are not issued as ServKeys*)
       
   576 Goalw [KeyCryptKey_def]
       
   577  "evs : kerberos ==> ~ KeyCryptKey K (shrK A) evs";
       
   578 by (parts_induct_tac 1);
       
   579 qed "shrK_not_KeyCryptKey";
       
   580 
       
   581 (*The Tgs message associates ServKey with AuthKey and therefore not with any 
       
   582   other key AuthKey.*)
       
   583 Goalw [KeyCryptKey_def]
       
   584  "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |}) \
       
   585 \      : set evs;                                         \
       
   586 \    AuthKey' ~= AuthKey;  evs : kerberos |]                      \
       
   587 \ ==> ~ KeyCryptKey AuthKey' ServKey evs";
       
   588 by (blast_tac (claset() addDs [unique_ServKeys]) 1);
       
   589 qed "Says_Tgs_KeyCryptKey";
       
   590 
       
   591 Goal "[| KeyCryptKey AuthKey ServKey evs;  evs : kerberos |] \
       
   592 \     ==> ~ KeyCryptKey ServKey K evs";
       
   593 by (etac rev_mp 1);
       
   594 by (parts_induct_tac 1);
       
   595 by (Step_tac 1);
       
   596 by (ALLGOALS Asm_full_simp_tac);
       
   597 (*K4 splits into subcases*)
       
   598 by (blast_tac (claset() addEs spies_partsEs 
       
   599                        addSDs [AuthKey_not_KeyCryptKey]) 4);
       
   600 (*ServKey is fresh and so could not have been used, by new_keys_not_used*)
       
   601 by (force_tac (claset() addSDs [Says_imp_spies RS parts.Inj,
       
   602 				Crypt_imp_invKey_keysFor],
       
   603                       simpset() addsimps [KeyCryptKey_def]) 2); 
       
   604 (*Others by freshness*)
       
   605 by (REPEAT (blast_tac (claset() addSEs spies_partsEs) 1));
       
   606 qed "KeyCryptKey_not_KeyCryptKey";
       
   607 
       
   608 (*The only session keys that can be found with the help of session keys are
       
   609   those sent by Tgs in step K4.  *)
       
   610 
       
   611 (*We take some pains to express the property
       
   612   as a logical equivalence so that the simplifier can apply it.*)
       
   613 Goal "P --> (Key K : analz (Key``KK Un H)) --> (K:KK | Key K : analz H)  \
       
   614 \     ==>       \
       
   615 \     P --> (Key K : analz (Key``KK Un H)) = (K:KK | Key K : analz H)";
       
   616 by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1);
       
   617 qed "Key_analz_image_Key_lemma";
       
   618 
       
   619 Goal "[| KeyCryptKey K K' evs; evs : kerberos |] \
       
   620 \     ==> Key K' : analz (insert (Key K) (spies evs))";
       
   621 by (full_simp_tac (simpset() addsimps [KeyCryptKey_def]) 1);
       
   622 by (Clarify_tac 1);
       
   623 by (dresolve_tac [Says_imp_spies RS analz.Inj RS analz_insertI] 1);
       
   624 by Auto_tac;
       
   625 qed "KeyCryptKey_analz_insert";
       
   626 
       
   627 Goal "[| K : AuthKeys evs Un range shrK;  evs : kerberos |]  \
       
   628 \     ==> ALL SK. ~ KeyCryptKey SK K evs";
       
   629 by (asm_full_simp_tac (simpset() addsimps [KeyCryptKey_def]) 1);
       
   630 by (blast_tac (claset() addDs [Says_Tgs_message_form]) 1);
       
   631 qed "AuthKeys_are_not_KeyCryptKey";
       
   632 
       
   633 Goal "[| K ~: AuthKeys evs; \
       
   634 \        K ~: range shrK; evs : kerberos |]  \
       
   635 \     ==> ALL SK. ~ KeyCryptKey K SK evs";
       
   636 by (asm_full_simp_tac (simpset() addsimps [KeyCryptKey_def]) 1);
       
   637 by (blast_tac (claset() addDs [Says_Tgs_message_form]) 1);
       
   638 qed "not_AuthKeys_not_KeyCryptKey";
       
   639 
       
   640 
       
   641 (*****************SECRECY THEOREMS****************)
       
   642 
       
   643 (*For proofs involving analz.*)
       
   644 val analz_sees_tac = 
       
   645   EVERY 
       
   646    [REPEAT (FIRSTGOAL analz_mono_contra_tac),
       
   647     forward_tac [Oops_range_spies2] 10, 
       
   648     forward_tac [Oops_range_spies1] 9, 
       
   649     forward_tac [Says_tgs_message_form] 7,
       
   650     forward_tac [Says_kas_message_form] 5, 
       
   651     REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE, exE]
       
   652 		  ORELSE' hyp_subst_tac)];
       
   653 
       
   654 Goal "[| KK <= -(range shrK); Key K : analz (spies evs); evs: kerberos |]   \
       
   655 \     ==> Key K : analz (Key `` KK Un spies evs)";
       
   656 by (blast_tac (claset() addDs [impOfSubs analz_mono]) 1);
       
   657 qed "analz_mono_KK";
       
   658 
       
   659 (* Big simplification law for keys SK that are not crypted by keys in KK   *)
       
   660 (* It helps prove three, otherwise hard, facts about keys. These facts are *)
       
   661 (* exploited as simplification laws for analz, and also "limit the damage" *)
       
   662 (* in case of loss of a key to the spy. See ESORICS98.                     *)
       
   663 Goal "evs : kerberos ==>                                         \
       
   664 \     (ALL SK KK. KK <= -(range shrK) -->                   \
       
   665 \     (ALL K: KK. ~ KeyCryptKey K SK evs)   -->           \
       
   666 \     (Key SK : analz (Key``KK Un (spies evs))) =        \
       
   667 \     (SK : KK | Key SK : analz (spies evs)))";
       
   668 by (etac kerberos.induct 1);
       
   669 by analz_sees_tac;
       
   670 by (REPEAT_FIRST (rtac allI));
       
   671 by (REPEAT_FIRST (rtac (Key_analz_image_Key_lemma RS impI)));
       
   672 by (ALLGOALS  
       
   673     (asm_simp_tac 
       
   674      (analz_image_freshK_ss addsimps
       
   675         [KeyCryptKey_Says, shrK_not_KeyCryptKey, 
       
   676 	 Auth_fresh_not_KeyCryptKey, Serv_fresh_not_KeyCryptKey, 
       
   677 	 Says_Tgs_KeyCryptKey, Spy_analz_shrK])));
       
   678 (*Fake*) 
       
   679 by (spy_analz_tac 1);
       
   680 (* Base + K2 done by the simplifier! *)
       
   681 (*K3*)
       
   682 by (Blast_tac 1);
       
   683 (*K4*)
       
   684 by (blast_tac (claset() addEs spies_partsEs 
       
   685                         addSDs [AuthKey_not_KeyCryptKey]) 1);
       
   686 (*K5*)
       
   687 by (rtac impI 1);
       
   688 by (case_tac "Key ServKey : analz (spies evs5)" 1);
       
   689 (*If ServKey is compromised then the result follows directly...*)
       
   690 by (asm_simp_tac 
       
   691      (simpset() addsimps [analz_insert_eq, 
       
   692 			 impOfSubs (Un_upper2 RS analz_mono)]) 1);
       
   693 (*...therefore ServKey is uncompromised.*)
       
   694 by (case_tac "KeyCryptKey ServKey SK evs5" 1);
       
   695 by (asm_simp_tac analz_image_freshK_ss 2);
       
   696 (*The KeyCryptKey ServKey SK evs5 case leads to a contradiction.*)
       
   697 by (blast_tac (claset() addSEs [ServKey_not_KeyCryptKey RSN(2, rev_notE)]
       
   698 		        addEs spies_partsEs delrules [allE, ballE]) 1);
       
   699 (** Level 14: Oops1 and Oops2 **)
       
   700 by (ALLGOALS Clarify_tac);
       
   701 (*Oops 2*)
       
   702 by (case_tac "Key ServKey : analz (spies evsO2)" 2);
       
   703 by (ALLGOALS Asm_full_simp_tac);
       
   704 by (forward_tac [analz_mono_KK] 2
       
   705     THEN assume_tac 2
       
   706     THEN assume_tac 2);
       
   707 by (forward_tac [analz_cut] 2 THEN assume_tac 2);
       
   708 by (blast_tac (claset() addDs [analz_cut, impOfSubs analz_mono]) 2);
       
   709 by (dres_inst_tac [("x","SK")] spec 2);
       
   710 by (dres_inst_tac [("x","insert ServKey KK")] spec 2);
       
   711 by (forward_tac [Says_Tgs_message_form] 2 THEN assume_tac 2);
       
   712 by (Clarify_tac 2);
       
   713 by (forward_tac [Says_imp_spies RS parts.Inj RS parts.Body 
       
   714                  RS parts.Snd RS parts.Snd RS parts.Snd] 2);
       
   715 by (Asm_full_simp_tac 2);
       
   716 by (blast_tac (claset() addSDs [ServKey_not_KeyCryptKey] 
       
   717                         addDs [impOfSubs analz_mono]) 2);
       
   718 (*Level 28: Oops 1*)
       
   719 by (dres_inst_tac [("x","SK")] spec 1);
       
   720 by (dres_inst_tac [("x","insert AuthKey KK")] spec 1);
       
   721 by (case_tac "KeyCryptKey AuthKey SK evsO1" 1);
       
   722 by (ALLGOALS Asm_full_simp_tac);
       
   723 by (blast_tac (claset() addSDs [KeyCryptKey_analz_insert]) 1);
       
   724 by (blast_tac (claset() addDs [impOfSubs analz_mono]) 1);
       
   725 qed_spec_mp "Key_analz_image_Key";
       
   726 
       
   727 
       
   728 (* First simplification law for analz: no session keys encrypt  *)
       
   729 (* authentication keys or shared keys.                          *)
       
   730 Goal "[| evs : kerberos;  K : (AuthKeys evs) Un range shrK;      \
       
   731 \        SesKey ~: range shrK |]                                 \
       
   732 \     ==> Key K : analz (insert (Key SesKey) (spies evs)) = \
       
   733 \         (K = SesKey | Key K : analz (spies evs))";
       
   734 by (forward_tac [AuthKeys_are_not_KeyCryptKey] 1 THEN assume_tac 1);
       
   735 by (asm_full_simp_tac (analz_image_freshK_ss addsimps [Key_analz_image_Key]) 1);
       
   736 qed "analz_insert_freshK1";
       
   737 
       
   738 
       
   739 (* Second simplification law for analz: no service keys encrypt *)
       
   740 (* any other keys.					        *)
       
   741 Goal "[| evs : kerberos;  ServKey ~: (AuthKeys evs); ServKey ~: range shrK|]\
       
   742 \     ==> Key K : analz (insert (Key ServKey) (spies evs)) = \
       
   743 \         (K = ServKey | Key K : analz (spies evs))";
       
   744 by (forward_tac [not_AuthKeys_not_KeyCryptKey] 1 
       
   745     THEN assume_tac 1
       
   746     THEN assume_tac 1);
       
   747 by (asm_full_simp_tac (analz_image_freshK_ss addsimps [Key_analz_image_Key]) 1);
       
   748 qed "analz_insert_freshK2";
       
   749 
       
   750 
       
   751 (* Third simplification law for analz: only one authentication key *)
       
   752 (* encrypts a certain service key.                                 *)
       
   753 Goal  
       
   754  "[| Says Tgs A    \
       
   755 \           (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
       
   756 \             : set evs;          \ 
       
   757 \           AuthKey ~= AuthKey'; AuthKey' ~: range shrK; evs : kerberos |]    \
       
   758 \       ==> Key ServKey : analz (insert (Key AuthKey') (spies evs)) =  \
       
   759 \               (ServKey = AuthKey' | Key ServKey : analz (spies evs))";
       
   760 by (dres_inst_tac [("AuthKey'","AuthKey'")] Says_Tgs_KeyCryptKey 1);
       
   761 by (Blast_tac 1);
       
   762 by (assume_tac 1);
       
   763 by (asm_full_simp_tac (analz_image_freshK_ss addsimps [Key_analz_image_Key]) 1);
       
   764 qed "analz_insert_freshK3";
       
   765 
       
   766 
       
   767 (*a weakness of the protocol*)
       
   768 Goal "[| Says Tgs A    \
       
   769 \             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
       
   770 \          : set evs;          \ 
       
   771 \        Key AuthKey : analz (spies evs); evs : kerberos |]    \
       
   772 \     ==> Key ServKey : analz (spies evs)";
       
   773 by (force_tac (claset() addDs [Says_imp_spies RS analz.Inj RS 
       
   774 			       analz.Decrypt RS analz.Fst],
       
   775 	       simpset()) 1);
       
   776 qed "AuthKey_compromises_ServKey";
       
   777 
       
   778 
       
   779 (********************** Guarantees for Kas *****************************)
       
   780 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, \
       
   781 \                     Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|}\
       
   782 \          : parts (spies evs); \
       
   783 \        Key ServKey ~: analz (spies evs);                          \
       
   784 \        B ~= Tgs; evs : kerberos |]                            \
       
   785 \     ==> ServKey ~: AuthKeys evs";
       
   786 by (etac rev_mp 1);
       
   787 by (etac rev_mp 1);
       
   788 by (asm_full_simp_tac (simpset() addsimps [AuthKeys_def]) 1);
       
   789 by (parts_induct_tac 1);
       
   790 by (Fake_parts_insert_tac 1);
       
   791 by (ALLGOALS (blast_tac (claset() addSEs spies_partsEs)));
       
   792 bind_thm ("ServKey_notin_AuthKeys", result() RSN (2, rev_notE));
       
   793 bind_thm ("ServKey_notin_AuthKeysD", result());
       
   794 
       
   795 
       
   796 (** If Spy sees the Authentication Key sent in msg K2, then 
       
   797     the Key has expired  **)
       
   798 Goal "[| A ~: bad;  evs : kerberos |]           \
       
   799 \     ==> Says Kas A                             \
       
   800 \              (Crypt (shrK A)                      \
       
   801 \                 {|Key AuthKey, Agent Tgs, Number Tk,     \
       
   802 \         Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
       
   803 \           : set evs -->                 \
       
   804 \         Key AuthKey : analz (spies evs) -->                       \
       
   805 \         ExpirAuth Tk evs";
       
   806 by (etac kerberos.induct 1);
       
   807 by analz_sees_tac;
       
   808 by (ALLGOALS 
       
   809     (asm_simp_tac 
       
   810      (simpset() addsimps ([Says_Kas_message_form, 
       
   811 			   analz_insert_eq, not_parts_not_analz, 
       
   812 			   analz_insert_freshK1] @ pushes))));
       
   813 (*Fake*) 
       
   814 by (spy_analz_tac 1);
       
   815 (*K2*)
       
   816 by (blast_tac (claset() addSEs spies_partsEs
       
   817             addIs [parts_insertI, impOfSubs analz_subset_parts, less_SucI]) 1);
       
   818 (*K4*)
       
   819 by (blast_tac (claset() addSEs spies_partsEs) 1);
       
   820 (*Level 8: K5*)
       
   821 by (blast_tac (claset() addEs [ServKey_notin_AuthKeys]
       
   822                         addDs [Says_Kas_message_form, 
       
   823 	       		       Says_imp_spies RS parts.Inj]
       
   824                         addIs [less_SucI]) 1);
       
   825 (*Oops1*)
       
   826 by (blast_tac (claset() addDs [unique_AuthKeys] addIs [less_SucI]) 1);
       
   827 (*Oops2*)
       
   828 by (blast_tac (claset() addDs [Says_Tgs_message_form,
       
   829                                Says_Kas_message_form]) 1);
       
   830 val lemma = result() RS mp RS mp RSN(1,rev_notE);
       
   831 
       
   832 
       
   833 Goal "[| Says Kas A                                             \
       
   834 \             (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})  \
       
   835 \          : set evs;                                \
       
   836 \        ~ ExpirAuth Tk evs;                         \
       
   837 \        A ~: bad;  evs : kerberos |]            \
       
   838 \     ==> Key AuthKey ~: analz (spies evs)";
       
   839 by (forward_tac [Says_Kas_message_form] 1 THEN assume_tac 1);
       
   840 by (blast_tac (claset() addSDs [lemma]) 1);
       
   841 qed "Confidentiality_Kas";
       
   842 
       
   843 
       
   844 
       
   845 
       
   846 
       
   847 
       
   848 (********************** Guarantees for Tgs *****************************)
       
   849 
       
   850 (** If Spy sees the Service Key sent in msg K4, then 
       
   851     the Key has expired  **)
       
   852 Goal "[| A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]           \
       
   853 \  ==> Key AuthKey ~: analz (spies evs) --> \
       
   854 \      Says Tgs A            \
       
   855 \        (Crypt AuthKey                      \
       
   856 \           {|Key ServKey, Agent B, Number Tt,     \
       
   857 \             Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})\
       
   858 \        : set evs -->                 \
       
   859 \      Key ServKey : analz (spies evs) -->                       \
       
   860 \      ExpirServ Tt evs";
       
   861 by (etac kerberos.induct 1);
       
   862 (*The Oops1 case is unusual: must simplify Authkey ~: analz (spies (ev#evs))
       
   863   rather than weakening it to Authkey ~: analz (spies evs), for we then
       
   864   conclude AuthKey ~= AuthKeya.*)
       
   865 by (Clarify_tac 9);
       
   866 by analz_sees_tac;
       
   867 by (rotate_tac ~1 11);
       
   868 by (ALLGOALS 
       
   869     (asm_full_simp_tac 
       
   870      (simpset() addsimps ([Says_Kas_message_form, Says_Tgs_message_form,
       
   871 			   analz_insert_eq, not_parts_not_analz, 
       
   872 			   analz_insert_freshK1, analz_insert_freshK2] @ pushes))));
       
   873 (*Fake*) 
       
   874 by (spy_analz_tac 1);
       
   875 (*K2*)
       
   876 by (blast_tac (claset() addSEs spies_partsEs
       
   877             addIs [parts_insertI, impOfSubs analz_subset_parts, less_SucI]) 1);
       
   878 (*K4*)
       
   879 by (case_tac "A ~= Aa" 1);
       
   880 by (blast_tac (claset() addSEs spies_partsEs
       
   881                         addIs [less_SucI]) 1);
       
   882 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj RS parts.Fst, 
       
   883                                A_trusts_AuthTicket, 
       
   884                                Confidentiality_Kas,
       
   885                                impOfSubs analz_subset_parts]) 1);
       
   886 by (ALLGOALS Clarify_tac);
       
   887 (*Oops2*)
       
   888 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj, 
       
   889                                Key_unique_SesKey] addIs [less_SucI]) 3);
       
   890 (** Level 12 **)
       
   891 (*Oops1*)
       
   892 by (forward_tac [Says_Kas_message_form] 2);
       
   893 by (assume_tac 2);
       
   894 by (blast_tac (claset() addDs [analz_insert_freshK3,
       
   895 			       Says_Kas_message_form, Says_Tgs_message_form] 
       
   896                         addIs  [less_SucI]) 2);
       
   897 (** Level 16 **)
       
   898 by (thin_tac "Says Aa Tgs ?X : set ?evs" 1);
       
   899 by (forward_tac [Says_imp_spies RS parts.Inj RS ServKey_notin_AuthKeysD] 1);
       
   900 by (assume_tac 1 THEN Blast_tac 1 THEN assume_tac 1);
       
   901 by (rotate_tac ~1 1);
       
   902 by (asm_full_simp_tac (simpset() addsimps [analz_insert_freshK2]) 1);
       
   903 by (etac disjE 1);
       
   904 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj, 
       
   905                                Key_unique_SesKey]) 1);
       
   906 by (blast_tac (claset() addIs [less_SucI]) 1);
       
   907 val lemma = result() RS mp RS mp RS mp RSN(1,rev_notE);
       
   908 
       
   909 
       
   910 (* In the real world Tgs can't check wheter AuthKey is secure! *)
       
   911 Goal 
       
   912  "[| Says Tgs A      \
       
   913 \             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
       
   914 \             : set evs;              \
       
   915 \           Key AuthKey ~: analz (spies evs);        \
       
   916 \           ~ ExpirServ Tt evs;                         \
       
   917 \           A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
   918 \        ==> Key ServKey ~: analz (spies evs)";
       
   919 by (forward_tac [Says_Tgs_message_form] 1 THEN assume_tac 1);
       
   920 by (blast_tac (claset() addDs [lemma]) 1);
       
   921 qed "Confidentiality_Tgs1";
       
   922 
       
   923 (* In the real world Tgs CAN check what Kas sends! *)
       
   924 Goal 
       
   925  "[| Says Kas A                                             \
       
   926 \              (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})  \
       
   927 \             : set evs;                                \
       
   928 \           Says Tgs A      \
       
   929 \             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
       
   930 \             : set evs;              \
       
   931 \           ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
       
   932 \           A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
   933 \        ==> Key ServKey ~: analz (spies evs)";
       
   934 by (blast_tac (claset() addSDs [Confidentiality_Kas,
       
   935                                 Confidentiality_Tgs1]) 1);
       
   936 qed "Confidentiality_Tgs2";
       
   937 
       
   938 (*Most general form*)
       
   939 val Confidentiality_Tgs3 = A_trusts_AuthTicket RS Confidentiality_Tgs2;
       
   940 
       
   941 
       
   942 (********************** Guarantees for Alice *****************************)
       
   943 
       
   944 val Confidentiality_Auth_A = A_trusts_AuthKey RS Confidentiality_Kas;
       
   945 
       
   946 Goal
       
   947  "[| Says Kas A \
       
   948 \      (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) : set evs;\
       
   949 \    Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}     \
       
   950 \      : parts (spies evs);                                       \
       
   951 \    Key AuthKey ~: analz (spies evs);            \
       
   952 \    evs : kerberos |]         \
       
   953 \==> Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})\
       
   954 \      : set evs";
       
   955 by (forward_tac [Says_Kas_message_form] 1 THEN assume_tac 1);
       
   956 by (etac rev_mp 1);
       
   957 by (etac rev_mp 1);
       
   958 by (etac rev_mp 1);
       
   959 by (parts_induct_tac 1);
       
   960 by (Fake_parts_insert_tac 1);
       
   961 (*K2 and K4 remain*)
       
   962 by (blast_tac (claset() addEs spies_partsEs addSEs [MPair_parts] 
       
   963                         addSDs [unique_CryptKey]) 2);
       
   964 by (blast_tac (claset() addSDs [A_trusts_K4, Says_Tgs_message_form, 
       
   965 				AuthKeys_used]) 1);
       
   966 qed "A_trusts_K4_bis";
       
   967 
       
   968 
       
   969 Goal "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}  \
       
   970 \          : parts (spies evs);                              \
       
   971 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}     \
       
   972 \          : parts (spies evs);                                       \
       
   973 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
       
   974 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
   975 \     ==> Key ServKey ~: analz (spies evs)";
       
   976 by (dtac A_trusts_AuthKey 1);
       
   977 by (assume_tac 1);
       
   978 by (assume_tac 1);
       
   979 by (blast_tac (claset() addDs [Confidentiality_Kas, 
       
   980                                Says_Kas_message_form,
       
   981                                A_trusts_K4_bis, Confidentiality_Tgs2]) 1);
       
   982 qed "Confidentiality_Serv_A";
       
   983 
       
   984 
       
   985 (********************** Guarantees for Bob *****************************)
       
   986 (* Theorems for the refined model have suffix "refined"                *)
       
   987 
       
   988 Goal
       
   989 "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})\
       
   990 \            : set evs; evs : kerberos|]  \
       
   991 \  ==> EX Tk. Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
       
   992 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
       
   993 \            : set evs";
       
   994 by (etac rev_mp 1);
       
   995 by (parts_induct_tac 1);
       
   996 by Auto_tac;
       
   997 by (blast_tac (claset() addSDs [Says_imp_spies RS parts.Inj RS parts.Fst RS
       
   998                                A_trusts_AuthTicket]) 1);
       
   999 qed "K4_imp_K2";
       
  1000 
       
  1001 Goal
       
  1002 "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})\
       
  1003 \     : set evs; evs : kerberos|]  \
       
  1004 \  ==> EX Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
       
  1005 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
       
  1006 \            : set evs   \
       
  1007 \         & ServLife + Tt <= AuthLife + Tk)";
       
  1008 by (etac rev_mp 1);
       
  1009 by (parts_induct_tac 1);
       
  1010 by Auto_tac;
       
  1011 by (blast_tac (claset() addSDs [Says_imp_spies RS parts.Inj RS parts.Fst RS
       
  1012                                A_trusts_AuthTicket]) 1);
       
  1013 qed "K4_imp_K2_refined";
       
  1014 
       
  1015 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}  \
       
  1016 \          : parts (spies evs);  B ~= Tgs;  B ~: bad;       \
       
  1017 \        evs : kerberos |]                        \
       
  1018 \==> EX AuthKey. \
       
  1019 \      Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt,  \
       
  1020 \                  Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|}) \
       
  1021 \      : set evs";
       
  1022 by (etac rev_mp 1);
       
  1023 by (parts_induct_tac 1);
       
  1024 by (Fake_parts_insert_tac 1);
       
  1025 by (Blast_tac 1);
       
  1026 qed "B_trusts_ServKey";
       
  1027 
       
  1028 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
       
  1029 \          : parts (spies evs);  B ~= Tgs;  B ~: bad;       \
       
  1030 \        evs : kerberos |]                        \
       
  1031 \  ==> EX AuthKey Tk. Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
       
  1032 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
       
  1033 \            : set evs";
       
  1034 by (blast_tac (claset() addSDs [B_trusts_ServKey, K4_imp_K2]) 1);
       
  1035 qed "B_trusts_ServTicket_Kas";
       
  1036 
       
  1037 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
       
  1038 \          : parts (spies evs); B ~= Tgs; B ~: bad;       \
       
  1039 \        evs : kerberos |]                        \
       
  1040 \  ==> EX AuthKey Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
       
  1041 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
       
  1042 \            : set evs            \
       
  1043 \          & ServLife + Tt <= AuthLife + Tk)";
       
  1044 by (blast_tac (claset() addSDs [B_trusts_ServKey,K4_imp_K2_refined]) 1);
       
  1045 qed "B_trusts_ServTicket_Kas_refined";
       
  1046 
       
  1047 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
       
  1048 \          : parts (spies evs); B ~= Tgs; B ~: bad;        \
       
  1049 \        evs : kerberos |]                        \
       
  1050 \==> EX Tk AuthKey.        \
       
  1051 \    Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, \
       
  1052 \                  Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
       
  1053 \      : set evs         \ 
       
  1054 \    & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,  \
       
  1055 \                  Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|}) \
       
  1056 \      : set evs";
       
  1057 by (forward_tac [B_trusts_ServKey] 1);
       
  1058 by (etac exE 4);
       
  1059 by (forward_tac [K4_imp_K2] 4);
       
  1060 by (Blast_tac 5);
       
  1061 by (ALLGOALS assume_tac);
       
  1062 qed "B_trusts_ServTicket";
       
  1063 
       
  1064 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
       
  1065 \          : parts (spies evs); B ~= Tgs; B ~: bad;        \
       
  1066 \        evs : kerberos |]                        \
       
  1067 \==> EX Tk AuthKey.        \
       
  1068 \    (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, \
       
  1069 \                  Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
       
  1070 \      : set evs         \ 
       
  1071 \    & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,  \
       
  1072 \                  Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|}) \
       
  1073 \      : set evs         \
       
  1074 \    & ServLife + Tt <= AuthLife + Tk)";
       
  1075 by (forward_tac [B_trusts_ServKey] 1);
       
  1076 by (etac exE 4);
       
  1077 by (forward_tac [K4_imp_K2_refined] 4);
       
  1078 by (Blast_tac 5);
       
  1079 by (ALLGOALS assume_tac);
       
  1080 qed "B_trusts_ServTicket_refined";
       
  1081 
       
  1082 
       
  1083 Goal "[| ~ ExpirServ Tt evs; ServLife + Tt <= AuthLife + Tk |]        \
       
  1084 \  ==> ~ ExpirAuth Tk evs";
       
  1085 by (blast_tac (claset() addDs [leI,le_trans] addEs [leE]) 1);
       
  1086 qed "NotExpirServ_NotExpirAuth_refined";
       
  1087 
       
  1088 
       
  1089 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|} \
       
  1090 \          : parts (spies evs);                                        \
       
  1091 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|} \
       
  1092 \          : parts (spies evs);                                         \
       
  1093 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
       
  1094 \          : parts (spies evs);                     \
       
  1095 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;     \
       
  1096 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
  1097 \     ==> Key ServKey ~: analz (spies evs)";
       
  1098 by (forward_tac [A_trusts_AuthKey] 1);
       
  1099 by (forward_tac [Confidentiality_Kas] 3);
       
  1100 by (forward_tac [B_trusts_ServTicket] 6);
       
  1101 by (etac exE 9);
       
  1102 by (etac exE 9);
       
  1103 by (forward_tac [Says_Kas_message_form] 9);
       
  1104 by (blast_tac (claset() addDs [A_trusts_K4, 
       
  1105                                unique_ServKeys, unique_AuthKeys,
       
  1106                                Confidentiality_Tgs2]) 10);
       
  1107 by (ALLGOALS assume_tac);
       
  1108 (*
       
  1109 The proof above executes in 8 secs. It can be done in one command in 50 secs:
       
  1110 by (blast_tac (claset() addDs [A_trusts_AuthKey, A_trusts_K4,
       
  1111                                Says_Kas_message_form, B_trusts_ServTicket,
       
  1112                                unique_ServKeys, unique_AuthKeys,
       
  1113                                Confidentiality_Kas, 
       
  1114                                Confidentiality_Tgs2]) 1);
       
  1115 *)
       
  1116 qed "Confidentiality_B";
       
  1117 
       
  1118 
       
  1119 (*Most general form -- only for refined model! *)
       
  1120 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
       
  1121 \          : parts (spies evs);                      \
       
  1122 \        ~ ExpirServ Tt evs;                         \
       
  1123 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
  1124 \     ==> Key ServKey ~: analz (spies evs)";
       
  1125 by (blast_tac (claset() addDs [B_trusts_ServTicket_refined,
       
  1126 			       NotExpirServ_NotExpirAuth_refined, 
       
  1127                                Confidentiality_Tgs2]) 1);
       
  1128 qed "Confidentiality_B_refined";
       
  1129 
       
  1130 
       
  1131 (********************** Authenticity theorems *****************************)
       
  1132 
       
  1133 (***1. Session Keys authenticity: they originated with servers.***)
       
  1134 
       
  1135 (*Authenticity of AuthKey for A: "A_trusts_AuthKey"*)
       
  1136 
       
  1137 (*Authenticity of ServKey for A: "A_trusts_ServKey"*)
       
  1138 Goal "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|} \
       
  1139 \          : parts (spies evs);                                     \
       
  1140 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}   \
       
  1141 \          : parts (spies evs);                                        \
       
  1142 \        ~ ExpirAuth Tk evs; A ~: bad; evs : kerberos |]         \
       
  1143 \==>Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})\
       
  1144 \      : set evs";
       
  1145 by (forward_tac [A_trusts_AuthKey] 1 THEN assume_tac 1 THEN assume_tac 1);
       
  1146 by (blast_tac (claset() addDs [Confidentiality_Auth_A, A_trusts_K4_bis]) 1);
       
  1147 qed "A_trusts_ServKey"; 
       
  1148 (*Note: requires a temporal check*)
       
  1149 
       
  1150 
       
  1151 (*Authenticity of ServKey for B: "B_trusts_ServKey"*)
       
  1152 
       
  1153 (***2. Parties authenticity: each party verifies "the identity of
       
  1154        another party who generated some data" (quoted from Neuman & Ts'o).***)
       
  1155        
       
  1156        (*These guarantees don't assess whether two parties agree on 
       
  1157          the same session key: sending a message containing a key
       
  1158          doesn't a priori state knowledge of the key.***)
       
  1159 
       
  1160 (*B checks authenticity of A: theorems "A_Authenticity", 
       
  1161                                        "A_authenticity_refined" *)
       
  1162 Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);  \
       
  1163 \        Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, \
       
  1164 \                                    ServTicket|}) : set evs;       \
       
  1165 \        Key ServKey ~: analz (spies evs);                \
       
  1166 \        A ~: bad; B ~: bad; evs : kerberos |]   \
       
  1167 \==> Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} : set evs";
       
  1168 by (etac rev_mp 1);
       
  1169 by (etac rev_mp 1);
       
  1170 by (etac rev_mp 1);
       
  1171 by (etac kerberos.induct 1);
       
  1172 by (forward_tac [Says_ticket_in_parts_spies] 5);
       
  1173 by (forward_tac [Says_ticket_in_parts_spies] 7);
       
  1174 by (REPEAT (FIRSTGOAL analz_mono_contra_tac));
       
  1175 by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib])));
       
  1176 by (Fake_parts_insert_tac 1);
       
  1177 (*K3*)
       
  1178 by (blast_tac (claset() addEs spies_partsEs
       
  1179                         addDs [A_trusts_AuthKey,
       
  1180                                Says_Kas_message_form, 
       
  1181                                Says_Tgs_message_form]) 1);
       
  1182 (*K4*)
       
  1183 by (force_tac (claset() addSDs [Crypt_imp_keysFor], simpset()) 1); 
       
  1184 (*K5*)
       
  1185 by (blast_tac (claset() addDs [Key_unique_SesKey] addEs spies_partsEs) 1);
       
  1186 qed "Says_Auth";
       
  1187 
       
  1188 (*The second assumption tells B what kind of key ServKey is.*)
       
  1189 Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);     \
       
  1190 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
       
  1191 \          : parts (spies evs);                                         \
       
  1192 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}  \ 
       
  1193 \          : parts (spies evs);                                          \
       
  1194 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}  \
       
  1195 \          : parts (spies evs);                                            \
       
  1196 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;  \
       
  1197 \        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
       
  1198 \  ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},\
       
  1199 \                 Crypt ServKey {|Agent A, Number Ta|} |} : set evs";
       
  1200 by (forward_tac [Confidentiality_B] 1);
       
  1201 by (forward_tac [B_trusts_ServKey] 9);
       
  1202 by (etac exE 12);
       
  1203 by (blast_tac (claset() addSDs [Says_imp_spies RS parts.Inj, Key_unique_SesKey]
       
  1204                         addSIs [Says_Auth]) 12);
       
  1205 by (ALLGOALS assume_tac);
       
  1206 qed "A_Authenticity";
       
  1207 
       
  1208 (*Stronger form in the refined model*)
       
  1209 Goal "[| Crypt ServKey {|Agent A, Number Ta2|} : parts (spies evs);     \
       
  1210 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
       
  1211 \          : parts (spies evs);                                         \
       
  1212 \        ~ ExpirServ Tt evs;                                        \
       
  1213 \        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
       
  1214 \  ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},\
       
  1215 \                 Crypt ServKey {|Agent A, Number Ta2|} |} : set evs";
       
  1216 by (forward_tac [Confidentiality_B_refined] 1);
       
  1217 by (forward_tac [B_trusts_ServKey] 6);
       
  1218 by (etac exE 9);
       
  1219 by (blast_tac (claset() addSDs [Says_imp_spies RS parts.Inj, Key_unique_SesKey]
       
  1220                         addSIs [Says_Auth]) 9);
       
  1221 by (ALLGOALS assume_tac);
       
  1222 qed "A_Authenticity_refined";
       
  1223 
       
  1224 
       
  1225 (*A checks authenticity of B: theorem "B_authenticity"*)
       
  1226 
       
  1227 Goal "[| Crypt ServKey (Number Ta) : parts (spies evs);  \
       
  1228 \        Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, \
       
  1229 \                                    ServTicket|}) : set evs;       \
       
  1230 \        Key ServKey ~: analz (spies evs);                \
       
  1231 \        A ~: bad; B ~: bad; evs : kerberos |]   \
       
  1232 \     ==> Says B A (Crypt ServKey (Number Ta)) : set evs";
       
  1233 by (etac rev_mp 1);
       
  1234 by (etac rev_mp 1);
       
  1235 by (etac rev_mp 1);
       
  1236 by (etac kerberos.induct 1);
       
  1237 by (forward_tac [Says_ticket_in_parts_spies] 5);
       
  1238 by (forward_tac [Says_ticket_in_parts_spies] 7);
       
  1239 by (REPEAT (FIRSTGOAL analz_mono_contra_tac));
       
  1240 by (ALLGOALS Asm_simp_tac);
       
  1241 by (Fake_parts_insert_tac 1);
       
  1242 by (force_tac (claset() addSDs [Crypt_imp_keysFor], simpset()) 1); 
       
  1243 by (Clarify_tac 1);
       
  1244 by (forward_tac [Says_Tgs_message_form] 1 THEN assume_tac 1);
       
  1245 by (Clarify_tac 1);  (*PROOF FAILED if omitted*)
       
  1246 by (blast_tac (claset() addDs [unique_CryptKey] addEs spies_partsEs) 1);
       
  1247 qed "Says_K6";
       
  1248 
       
  1249 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|}   \
       
  1250 \          : parts (spies evs);    \
       
  1251 \        Key AuthKey ~: analz (spies evs); AuthKey ~: range shrK;  \
       
  1252 \        evs : kerberos |]              \
       
  1253 \ ==> EX A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|})\
       
  1254 \             : set evs";
       
  1255 by (etac rev_mp 1);
       
  1256 by (etac rev_mp 1);
       
  1257 by (parts_induct_tac 1);
       
  1258 by (Fake_parts_insert_tac 1);
       
  1259 by (Blast_tac 1);
       
  1260 by (Blast_tac 1);
       
  1261 qed "K4_trustworthy";
       
  1262 
       
  1263 Goal "[| Crypt ServKey (Number Ta) : parts (spies evs);           \
       
  1264 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|} \
       
  1265 \          : parts (spies evs);                                        \ 
       
  1266 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
       
  1267 \          : parts (spies evs);                                          \
       
  1268 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
       
  1269 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
  1270 \     ==> Says B A (Crypt ServKey (Number Ta)) : set evs";
       
  1271 by (forward_tac [A_trusts_AuthKey] 1);
       
  1272 by (forward_tac [Says_Kas_message_form] 3);
       
  1273 by (forward_tac [Confidentiality_Kas] 4);
       
  1274 by (forward_tac [K4_trustworthy] 7);
       
  1275 by (Blast_tac 8);
       
  1276 by (etac exE 9);
       
  1277 by (forward_tac [K4_imp_K2] 9);
       
  1278 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj, Key_unique_SesKey]
       
  1279                         addSIs [Says_K6]
       
  1280                         addSEs [Confidentiality_Tgs1 RSN (2,rev_notE)]) 10);
       
  1281 by (ALLGOALS assume_tac);
       
  1282 qed "B_Authenticity";
       
  1283 
       
  1284 
       
  1285 (***3. Parties' knowledge of session keys. A knows a session key if she
       
  1286        used it to build a cipher.***)
       
  1287 
       
  1288 Goal "[| Says B A (Crypt ServKey (Number Ta)) : set evs;           \
       
  1289 \        Key ServKey ~: analz (spies evs);                          \
       
  1290 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
  1291 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
       
  1292 by (simp_tac (simpset() addsimps [Issues_def]) 1);
       
  1293 by (rtac exI 1);
       
  1294 by (rtac conjI 1);
       
  1295 by (assume_tac 1);
       
  1296 by (Simp_tac 1);
       
  1297 by (etac rev_mp 1);
       
  1298 by (etac rev_mp 1);
       
  1299 by (etac kerberos.induct 1);
       
  1300 by (forward_tac [Says_ticket_in_parts_spies] 5);
       
  1301 by (forward_tac [Says_ticket_in_parts_spies] 7);
       
  1302 by (REPEAT (FIRSTGOAL analz_mono_contra_tac));
       
  1303 by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib])));
       
  1304 by (Fake_parts_insert_tac 1);
       
  1305 (*K6 requires numerous lemmas*)
       
  1306 by (asm_full_simp_tac (simpset() addsimps [takeWhile_tail]) 1);
       
  1307 by (blast_tac (claset() addDs [B_trusts_ServTicket,
       
  1308                                impOfSubs parts_spies_takeWhile_mono,
       
  1309                                impOfSubs parts_spies_evs_revD2]
       
  1310                         addIs [Says_K6]
       
  1311                         addEs spies_partsEs) 1);
       
  1312 qed "B_Knows_B_Knows_ServKey_lemma";
       
  1313 (*Key ServKey ~: analz (spies evs) could be relaxed by Confidentiality_B
       
  1314   but this is irrelevant because B knows what he knows!                  *)
       
  1315 
       
  1316 Goal "[| Says B A (Crypt ServKey (Number Ta)) : set evs;           \
       
  1317 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}\
       
  1318 \           : parts (spies evs);\
       
  1319 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}\
       
  1320 \           : parts (spies evs);\
       
  1321 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
       
  1322 \          : parts (spies evs);     \
       
  1323 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;              \
       
  1324 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
  1325 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
       
  1326 by (blast_tac (claset() addSDs [Confidentiality_B,
       
  1327 	                       B_Knows_B_Knows_ServKey_lemma]) 1);
       
  1328 qed "B_Knows_B_Knows_ServKey";
       
  1329 
       
  1330 Goal "[| Says B A (Crypt ServKey (Number Ta)) : set evs;           \
       
  1331 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}\
       
  1332 \           : parts (spies evs);\
       
  1333 \        ~ ExpirServ Tt evs;            \
       
  1334 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
  1335 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
       
  1336 by (blast_tac (claset() addSDs [Confidentiality_B_refined,
       
  1337 	                       B_Knows_B_Knows_ServKey_lemma]) 1);
       
  1338 qed "B_Knows_B_Knows_ServKey_refined";
       
  1339 
       
  1340 
       
  1341 Goal "[| Crypt ServKey (Number Ta) : parts (spies evs);           \
       
  1342 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|} \
       
  1343 \          : parts (spies evs);                                        \ 
       
  1344 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
       
  1345 \          : parts (spies evs);                                          \
       
  1346 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
       
  1347 \        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
       
  1348 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
       
  1349 by (blast_tac (claset() addSDs [B_Authenticity, Confidentiality_Serv_A,
       
  1350                                 B_Knows_B_Knows_ServKey_lemma]) 1);
       
  1351 qed "A_Knows_B_Knows_ServKey";
       
  1352 
       
  1353 Goal "[| Says A Tgs     \
       
  1354 \            {|AuthTicket, Crypt AuthKey {|Agent A, Number Ta|}, Agent B|}\
       
  1355 \          : set evs;      \
       
  1356 \        A ~: bad;  evs : kerberos |]         \
       
  1357 \     ==> EX Tk. Says Kas A (Crypt (shrK A) \
       
  1358 \                     {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \
       
  1359 \                  : set evs";
       
  1360 by (etac rev_mp 1);
       
  1361 by (parts_induct_tac 1);
       
  1362 by (Fake_parts_insert_tac 1);
       
  1363 by (Blast_tac 1);
       
  1364 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj RS 
       
  1365 			       A_trusts_AuthKey]) 1);
       
  1366 qed "K3_imp_K2";
       
  1367 
       
  1368 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}   \
       
  1369 \          : parts (spies evs);                    \
       
  1370 \        Says Kas A (Crypt (shrK A) \
       
  1371 \                    {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \
       
  1372 \        : set evs;    \
       
  1373 \        Key AuthKey ~: analz (spies evs);       \
       
  1374 \        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
       
  1375 \  ==> Says Tgs A (Crypt AuthKey        \ 
       
  1376 \                    {|Key ServKey, Agent B, Number Tt, ServTicket|})  \
       
  1377 \        : set evs";      
       
  1378 by (etac rev_mp 1);
       
  1379 by (etac rev_mp 1);
       
  1380 by (etac rev_mp 1);
       
  1381 by (parts_induct_tac 1);
       
  1382 by (Fake_parts_insert_tac 1);
       
  1383 by (force_tac (claset() addSDs [Crypt_imp_keysFor], simpset()) 1); 
       
  1384 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj RS parts.Fst RS
       
  1385                                A_trusts_AuthTicket, unique_AuthKeys]) 1);
       
  1386 qed "K4_trustworthy'";
       
  1387 
       
  1388 Goal "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} \
       
  1389 \          : set evs;       \
       
  1390 \        Key ServKey ~: analz (spies evs);       \
       
  1391 \        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
       
  1392 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
       
  1393 by (simp_tac (simpset() addsimps [Issues_def]) 1);
       
  1394 by (rtac exI 1);
       
  1395 by (rtac conjI 1);
       
  1396 by (assume_tac 1);
       
  1397 by (Simp_tac 1);
       
  1398 by (etac rev_mp 1);
       
  1399 by (etac rev_mp 1);
       
  1400 by (etac kerberos.induct 1);
       
  1401 by (forward_tac [Says_ticket_in_parts_spies] 5);
       
  1402 by (forward_tac [Says_ticket_in_parts_spies] 7);
       
  1403 by (REPEAT (FIRSTGOAL analz_mono_contra_tac));
       
  1404 by (ALLGOALS Asm_simp_tac);
       
  1405 by (Clarify_tac 1);
       
  1406 (*K6*)
       
  1407 by Auto_tac;
       
  1408 by (asm_full_simp_tac (simpset() addsimps [takeWhile_tail]) 1);
       
  1409 (*Level 15: case study necessary because the assumption doesn't state
       
  1410   the form of ServTicket. The guarantee becomes stronger.*)
       
  1411 by (case_tac "Key AuthKey : analz (spies evs5)" 1);
       
  1412 by (force_tac (claset() addDs [Says_imp_spies RS analz.Inj RS 
       
  1413 			       analz.Decrypt RS analz.Fst],
       
  1414 	       simpset()) 1);
       
  1415 by (blast_tac (claset() addDs [K3_imp_K2, K4_trustworthy',
       
  1416                                impOfSubs parts_spies_takeWhile_mono,
       
  1417                                impOfSubs parts_spies_evs_revD2]
       
  1418                         addIs [Says_Auth] 
       
  1419                         addEs spies_partsEs) 1);
       
  1420 by (asm_full_simp_tac (simpset() addsimps [takeWhile_tail]) 1);
       
  1421 qed "A_Knows_A_Knows_ServKey_lemma";
       
  1422 
       
  1423 Goal "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} \
       
  1424 \          : set evs;       \
       
  1425 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
       
  1426 \          : parts (spies evs);\
       
  1427 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}\
       
  1428 \          : parts (spies evs);                                        \
       
  1429 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;\
       
  1430 \        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
       
  1431 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
       
  1432 by (blast_tac (claset() addSDs [Confidentiality_Serv_A,
       
  1433 	                       A_Knows_A_Knows_ServKey_lemma]) 1);
       
  1434 qed "A_Knows_A_Knows_ServKey";
       
  1435 
       
  1436 Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);     \
       
  1437 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
       
  1438 \          : parts (spies evs);                                         \
       
  1439 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}  \ 
       
  1440 \          : parts (spies evs);                                          \
       
  1441 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}  \
       
  1442 \          : parts (spies evs);                                            \
       
  1443 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;  \
       
  1444 \        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
       
  1445 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
       
  1446 by (blast_tac (claset() addDs [A_Authenticity, Confidentiality_B,
       
  1447 	                       A_Knows_A_Knows_ServKey_lemma]) 1);
       
  1448 qed "B_Knows_A_Knows_ServKey";
       
  1449 
       
  1450 
       
  1451 Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);     \
       
  1452 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
       
  1453 \          : parts (spies evs);                                         \
       
  1454 \        ~ ExpirServ Tt evs;                                        \
       
  1455 \        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
       
  1456 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
       
  1457 by (blast_tac (claset() addDs [A_Authenticity_refined, 
       
  1458                                Confidentiality_B_refined,
       
  1459 	                       A_Knows_A_Knows_ServKey_lemma]) 1);
       
  1460 qed "B_Knows_A_Knows_ServKey_refined";