Theory Kerberos_BAN_Gets

(*  Title:      HOL/Auth/Kerberos_BAN_Gets.thy
    Author:     Giampaolo Bella, Catania University
*)

section‹The Kerberos Protocol, BAN Version, with Gets event›

theory Kerberos_BAN_Gets imports Public begin

text‹From page 251 of
  Burrows, Abadi and Needham (1989).  A Logic of Authentication.
  Proc. Royal Soc. 426

  Confidentiality (secrecy) and authentication properties rely on
  temporal checks: strong guarantees in a little abstracted - but
  very realistic - model.
›

(* Temporal modelization: session keys can be leaked
                          ONLY when they have expired *)

consts

    (*Duration of the session key*)
    sesKlife   :: nat

    (*Duration of the authenticator*)
    authlife :: nat

text‹The ticket should remain fresh for two journeys on the network at least›
text‹The Gets event causes longer traces for the protocol to reach its end›
specification (sesKlife)
  sesKlife_LB [iff]: "4  sesKlife"
    by blast

text‹The authenticator only for one journey›
text‹The Gets event causes longer traces for the protocol to reach its end›
specification (authlife)
  authlife_LB [iff]:    "2  authlife"
    by blast


abbreviation
  CT :: "event list  nat" where
  "CT == length"

abbreviation
  expiredK :: "[nat, event list]  bool" where
  "expiredK T evs == sesKlife + T < CT evs"

abbreviation
  expiredA :: "[nat, event list]  bool" where
  "expiredA T evs == authlife + T < CT evs"


definition
 (* Yields the subtrace of a given trace from its beginning to a given event *)
  before :: "[event, event list]  event list" ("before _ on _")
  where "before ev on evs = takeWhile (λz. z  ev) (rev evs)"

definition
 (* States than an event really appears only once on a trace *)
  Unique :: "[event, event list]  bool" ("Unique _ on _")
  where "Unique ev on evs = (ev  set (tl (dropWhile (λz. z  ev) evs)))"


inductive_set bankerb_gets :: "event list set"
 where

   Nil:  "[]  bankerb_gets"

 | Fake: " evsf  bankerb_gets;  X  synth (analz (knows Spy evsf)) 
           Says Spy B X # evsf  bankerb_gets"

 | Reception: " evsr bankerb_gets; Says A B X  set evsr 
                 Gets B X # evsr  bankerb_gets"

 | BK1:  " evs1  bankerb_gets 
           Says A Server Agent A, Agent B # evs1
                  bankerb_gets"


 | BK2:  " evs2  bankerb_gets;  Key K  used evs2; K  symKeys;
             Gets Server Agent A, Agent B  set evs2 
           Says Server A
                (Crypt (shrK A)
                   Number (CT evs2), Agent B, Key K,
                    (Crypt (shrK B) Number (CT evs2), Agent A, Key K))
                # evs2  bankerb_gets"


 | BK3:  " evs3  bankerb_gets;
             Gets A (Crypt (shrK A) Number Tk, Agent B, Key K, Ticket)
                set evs3;
             Says A Server Agent A, Agent B  set evs3;
             ¬ expiredK Tk evs3 
           Says A B Ticket, Crypt K Agent A, Number (CT evs3) 
               # evs3  bankerb_gets"


 | BK4:  " evs4  bankerb_gets;
             Gets B (Crypt (shrK B) Number Tk, Agent A, Key K),
                         (Crypt K Agent A, Number Ta)   set evs4;
             ¬ expiredK Tk evs4;  ¬ expiredA Ta evs4 
           Says B A (Crypt K (Number Ta)) # evs4
                 bankerb_gets"

        (*Old session keys may become compromised*)
 | Oops: " evso  bankerb_gets;
         Says Server A (Crypt (shrK A) Number Tk, Agent B, Key K, Ticket)
                set evso;
             expiredK Tk evso 
           Notes Spy Number Tk, Key K # evso  bankerb_gets"


declare Says_imp_knows_Spy [THEN parts.Inj, dest]
declare parts.Body [dest]
declare analz_into_parts [dest]
declare Fake_parts_insert_in_Un [dest]
declare knows_Spy_partsEs [elim]


text‹A "possibility property": there are traces that reach the end.›
lemma "Key K  used []; K  symKeys
        Timestamp. evs  bankerb_gets.
             Says B A (Crypt K (Number Timestamp))
                   set evs"
apply (cut_tac sesKlife_LB)
apply (cut_tac authlife_LB)
apply (intro exI bexI)
apply (rule_tac [2]
           bankerb_gets.Nil [THEN bankerb_gets.BK1, THEN bankerb_gets.Reception,
                            THEN bankerb_gets.BK2, THEN bankerb_gets.Reception,
                            THEN bankerb_gets.BK3, THEN bankerb_gets.Reception,
                            THEN bankerb_gets.BK4])
apply (possibility, simp_all (no_asm_simp) add: used_Cons)
done


text‹Lemmas about reception invariant: if a message is received it certainly
was sent›
lemma Gets_imp_Says :
     " Gets B X  set evs; evs  bankerb_gets   A. Says A B X  set evs"
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply auto
done

lemma Gets_imp_knows_Spy: 
     " Gets B X  set evs; evs  bankerb_gets    X  knows Spy evs"
apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
done

lemma Gets_imp_knows_Spy_parts[dest]:
    " Gets B X  set evs; evs  bankerb_gets    X  parts (knows Spy evs)"
apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj])
done

lemma Gets_imp_knows:
     " Gets B X  set evs; evs  bankerb_gets    X  knows B evs"
by (metis Gets_imp_knows_Spy Gets_imp_knows_agents)

lemma Gets_imp_knows_analz:
    " Gets B X  set evs; evs  bankerb_gets    X  analz (knows B evs)"
apply (blast dest: Gets_imp_knows [THEN analz.Inj])
done

text‹Lemmas for reasoning about predicate "before"›
lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X}  (used evs)"
apply (induct_tac "evs")
apply simp
apply (rename_tac a b)
apply (induct_tac "a")
apply auto
done

lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X}  (used evs)"
apply (induct_tac "evs")
apply simp
apply (rename_tac a b)
apply (induct_tac "a")
apply auto
done

lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs"
apply (induct_tac "evs")
apply simp
apply (rename_tac a b)
apply (induct_tac "a")
apply auto
done

lemma used_evs_rev: "used evs = used (rev evs)"
apply (induct_tac "evs")
apply simp
apply (rename_tac a b)
apply (induct_tac "a")
apply (simp add: used_Says_rev)
apply (simp add: used_Gets_rev)
apply (simp add: used_Notes_rev)
done

lemma used_takeWhile_used [rule_format]: 
      "x  used (takeWhile P X)  x  used X"
apply (induct_tac "X")
apply simp
apply (rename_tac a b)
apply (induct_tac "a")
apply (simp_all add: used_Nil)
apply (blast dest!: initState_into_used)+
done

lemma set_evs_rev: "set evs = set (rev evs)"
apply auto
done

lemma takeWhile_void [rule_format]:
      "x  set evs  takeWhile (λz. z  x) evs = evs"
apply auto
done

(**** Inductive proofs about bankerb_gets ****)

text‹Forwarding Lemma for reasoning about the encrypted portion of message BK3›
lemma BK3_msg_in_parts_knows_Spy:
     "Gets A (Crypt KA Timestamp, B, K, X)  set evs; evs  bankerb_gets  
       X  parts (knows Spy evs)"
apply blast
done

lemma Oops_parts_knows_Spy:
     "Says Server A (Crypt (shrK A) Timestamp, B, K, X)  set evs
       K  parts (knows Spy evs)"
apply blast
done


text‹Spy never sees another agent's shared key! (unless it's bad at start)›
lemma Spy_see_shrK [simp]:
     "evs  bankerb_gets  (Key (shrK A)  parts (knows Spy evs)) = (A  bad)"
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Oops_parts_knows_Spy)
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast+)
done


lemma Spy_analz_shrK [simp]:
     "evs  bankerb_gets  (Key (shrK A)  analz (knows Spy evs)) = (A  bad)"
by auto

lemma Spy_see_shrK_D [dest!]:
     " Key (shrK A)  parts (knows Spy evs);
                evs  bankerb_gets   Abad"
by (blast dest: Spy_see_shrK)

lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D,  dest!]


text‹Nobody can have used non-existent keys!›
lemma new_keys_not_used [simp]:
    "Key K  used evs; K  symKeys; evs  bankerb_gets
      K  keysFor (parts (knows Spy evs))"
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Oops_parts_knows_Spy)
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)
txt‹Fake›
apply (force dest!: keysFor_parts_insert)
txt‹BK2, BK3, BK4›
apply (force dest!: analz_shrK_Decrypt)+
done

subsection‹Lemmas concerning the form of items passed in messages›

text‹Describes the form of K, X and K' when the Server sends this message.›
lemma Says_Server_message_form:
     " Says Server A (Crypt K' Number Tk, Agent B, Key K, Ticket)
          set evs; evs  bankerb_gets 
       K' = shrK A  K  range shrK 
          Ticket = (Crypt (shrK B) Number Tk, Agent A, Key K) 
          Key K  used(before
                  Says Server A (Crypt K' Number Tk, Agent B, Key K, Ticket)
                  on evs) 
          Tk = CT(before 
                  Says Server A (Crypt K' Number Tk, Agent B, Key K, Ticket)
                  on evs)"
unfolding before_def
apply (erule rev_mp)
apply (erule bankerb_gets.induct, simp_all)
txt‹We need this simplification only for Message 2›
apply (simp (no_asm) add: takeWhile_tail)
apply auto
txt‹Two subcases of Message 2. Subcase: used before›
apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] 
                   used_takeWhile_used)
txt‹subcase: CT before›
apply (fastforce dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])
done


text‹If the encrypted message appears then it originated with the Server
  PROVIDED that A is NOT compromised!
  This allows A to verify freshness of the session key.
›
lemma Kab_authentic:
     " Crypt (shrK A) Number Tk, Agent B, Key K, X
            parts (knows Spy evs);
         A  bad;  evs  bankerb_gets 
        Says Server A (Crypt (shrK A) Number Tk, Agent B, Key K, X)
              set evs"
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Oops_parts_knows_Spy)
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast)
done


text‹If the TICKET appears then it originated with the Server›
text‹FRESHNESS OF THE SESSION KEY to B›
lemma ticket_authentic:
     " Crypt (shrK B) Number Tk, Agent A, Key K  parts (knows Spy evs);
         B  bad;  evs  bankerb_gets 
        Says Server A
            (Crypt (shrK A) Number Tk, Agent B, Key K,
                          Crypt (shrK B) Number Tk, Agent A, Key K)
            set evs"
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Oops_parts_knows_Spy)
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all, blast)
done


text‹EITHER describes the form of X when the following message is sent,
  OR     reduces it to the Fake case.
  Use Says_Server_message_form› if applicable.›
lemma Gets_Server_message_form:
     " Gets A (Crypt (shrK A) Number Tk, Agent B, Key K, X)
             set evs;
         evs  bankerb_gets 
  (K  range shrK  X = (Crypt (shrK B) Number Tk, Agent A, Key K))
          | X  analz (knows Spy evs)"
apply (case_tac "A  bad")
apply (force dest!: Gets_imp_knows_Spy [THEN analz.Inj])
apply (blast dest!: Kab_authentic Says_Server_message_form)
done


text‹Reliability guarantees: honest agents act as we expect›

lemma BK3_imp_Gets:
   " Says A B Ticket, Crypt K Agent A, Number Ta  set evs;
      A  bad; evs  bankerb_gets 
      Tk. Gets A (Crypt (shrK A) Number Tk, Agent B, Key K, Ticket)
       set evs"
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply auto
done

lemma BK4_imp_Gets:
   " Says B A (Crypt K (Number Ta))  set evs;
      B  bad; evs  bankerb_gets 
    Tk. Gets B Crypt (shrK B) Number Tk, Agent A, Key K,
                    Crypt K Agent A, Number Ta  set evs"
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply auto
done

lemma Gets_A_knows_K:
  " Gets A (Crypt (shrK A) Number Tk, Agent B, Key K, Ticket)  set evs;
     evs  bankerb_gets 
  Key K  analz (knows A evs)"
apply (force dest: Gets_imp_knows_analz)
done

lemma Gets_B_knows_K:
  " Gets B Crypt (shrK B) Number Tk, Agent A, Key K,
             Crypt K Agent A, Number Ta  set evs;
     evs  bankerb_gets 
  Key K  analz (knows B evs)"
apply (force dest: Gets_imp_knows_analz)
done


(****
 The following is to prove theorems of the form

  Key K ∈ analz (insert (Key KAB) (knows Spy evs)) ⟹
  Key K ∈ analz (knows Spy evs)

 A more general formula must be proved inductively.

****)


text‹Session keys are not used to encrypt other session keys›
lemma analz_image_freshK [rule_format (no_asm)]:
     "evs  bankerb_gets 
   K KK. KK  - (range shrK) 
          (Key K  analz (Key`KK  (knows Spy evs))) =
          (K  KK | Key K  analz (knows Spy evs))"
apply (erule bankerb_gets.induct)
apply (drule_tac [8] Says_Server_message_form)
apply (erule_tac [6] Gets_Server_message_form [THEN disjE], analz_freshK, spy_analz, auto) 
done


lemma analz_insert_freshK:
     " evs  bankerb_gets;  KAB  range shrK  
      (Key K  analz (insert (Key KAB) (knows Spy evs))) =
      (K = KAB | Key K  analz (knows Spy evs))"
by (simp only: analz_image_freshK analz_image_freshK_simps)


text‹The session key K uniquely identifies the message›
lemma unique_session_keys:
     " Says Server A
           (Crypt (shrK A) Number Tk, Agent B, Key K, X)  set evs;
         Says Server A'
          (Crypt (shrK A') Number Tk', Agent B', Key K, X')  set evs;
         evs  bankerb_gets   A=A'  Tk=Tk'  B=B'  X = X'"
apply (erule rev_mp)
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Oops_parts_knows_Spy)
apply (frule_tac [6] BK3_msg_in_parts_knows_Spy, simp_all)
txt‹BK2: it can't be a new key›
apply blast
done

lemma unique_session_keys_Gets:
     " Gets A
           (Crypt (shrK A) Number Tk, Agent B, Key K, X)  set evs;
        Gets A
          (Crypt (shrK A) Number Tk', Agent B', Key K, X')  set evs;
        A  bad; evs  bankerb_gets   Tk=Tk'  B=B'  X = X'"
apply (blast dest!: Kab_authentic unique_session_keys)
done


lemma Server_Unique:
     " Says Server A
            (Crypt (shrK A) Number Tk, Agent B, Key K, Ticket)  set evs;
        evs  bankerb_gets   
   Unique Says Server A (Crypt (shrK A) Number Tk, Agent B, Key K, Ticket)
   on evs"
apply (erule rev_mp, erule bankerb_gets.induct, simp_all add: Unique_def)
apply blast
done



subsection‹Non-temporal guarantees, explicitly relying on non-occurrence of
oops events - refined below by temporal guarantees›

text‹Non temporal treatment of confidentiality›

text‹Lemma: the session key sent in msg BK2 would be lost by oops
    if the spy could see it!›
lemma lemma_conf [rule_format (no_asm)]:
     " A  bad;  B  bad;  evs  bankerb_gets 
   Says Server A
          (Crypt (shrK A) Number Tk, Agent B, Key K,
                            Crypt (shrK B) Number Tk, Agent A, Key K)
          set evs 
      Key K  analz (knows Spy evs)  Notes Spy Number Tk, Key K  set evs"
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Says_Server_message_form)
apply (frule_tac [6] Gets_Server_message_form [THEN disjE])
apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes)
txt‹Fake›
apply spy_analz
txt‹BK2›
apply (blast intro: parts_insertI)
txt‹BK3›
apply (case_tac "Aa  bad")
 prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz)
txt‹Oops›
apply (blast dest: unique_session_keys)
done


text‹Confidentiality for the Server: Spy does not see the keys sent in msg BK2
as long as they have not expired.›
lemma Confidentiality_S:
     " Says Server A
          (Crypt K' Number Tk, Agent B, Key K, Ticket)  set evs;
        Notes Spy Number Tk, Key K  set evs;
         A  bad;  B  bad;  evs  bankerb_gets
        Key K  analz (knows Spy evs)"
apply (frule Says_Server_message_form, assumption)
apply (blast intro: lemma_conf)
done

text‹Confidentiality for Alice›
lemma Confidentiality_A:
     " Crypt (shrK A) Number Tk, Agent B, Key K, X  parts (knows Spy evs);
        Notes Spy Number Tk, Key K  set evs;
        A  bad;  B  bad;  evs  bankerb_gets
        Key K  analz (knows Spy evs)"
by (blast dest!: Kab_authentic Confidentiality_S)

text‹Confidentiality for Bob›
lemma Confidentiality_B:
     " Crypt (shrK B) Number Tk, Agent A, Key K
           parts (knows Spy evs);
        Notes Spy Number Tk, Key K  set evs;
        A  bad;  B  bad;  evs  bankerb_gets
        Key K  analz (knows Spy evs)"
by (blast dest!: ticket_authentic Confidentiality_S)


text‹Non temporal treatment of authentication›

text‹Lemmas lemma_A› and lemma_B› in fact are common to both temporal and non-temporal treatments›
lemma lemma_A [rule_format]:
     " A  bad; B  bad; evs  bankerb_gets 
      
         Key K  analz (knows Spy evs) 
         Says Server A (Crypt (shrK A) Number Tk, Agent B, Key K, X)
          set evs 
          Crypt K Agent A, Number Ta  parts (knows Spy evs) 
         Says A B X, Crypt K Agent A, Number Ta
              set evs"
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Oops_parts_knows_Spy)
apply (frule_tac [6] Gets_Server_message_form)
apply (frule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)
apply (simp_all (no_asm_simp) add: all_conj_distrib)
txt‹Fake›
apply blast
txt‹BK2›
apply (force dest: Crypt_imp_invKey_keysFor)
txt‹BK3›
apply (blast dest: Kab_authentic unique_session_keys)
done
lemma lemma_B [rule_format]:
     " B  bad;  evs  bankerb_gets 
       Key K  analz (knows Spy evs) 
          Says Server A (Crypt (shrK A) Number Tk, Agent B, Key K, X)
           set evs 
          Crypt K (Number Ta)  parts (knows Spy evs) 
          Says B A (Crypt K (Number Ta))  set evs"
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Oops_parts_knows_Spy)
apply (frule_tac [6] Gets_Server_message_form)
apply (drule_tac [7] BK3_msg_in_parts_knows_Spy, analz_mono_contra)
apply (simp_all (no_asm_simp) add: all_conj_distrib)
txt‹Fake›
apply blast
txt‹BK2› 
apply (force dest: Crypt_imp_invKey_keysFor)
txt‹BK4›
apply (blast dest: ticket_authentic unique_session_keys
                   Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad)
done


text‹The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.›

text‹Authentication of A to B›
lemma B_authenticates_A_r:
     " Crypt K Agent A, Number Ta  parts (knows Spy evs);
         Crypt (shrK B) Number Tk, Agent A, Key K   parts (knows Spy evs);
        Notes Spy Number Tk, Key K  set evs;
         A  bad;  B  bad;  evs  bankerb_gets 
       Says A B Crypt (shrK B) Number Tk, Agent A, Key K,
                     Crypt K Agent A, Number Ta  set evs"
by (blast dest!: ticket_authentic
          intro!: lemma_A
          elim!: Confidentiality_S [THEN [2] rev_notE])

text‹Authentication of B to A›
lemma A_authenticates_B_r:
     " Crypt K (Number Ta)  parts (knows Spy evs);
        Crypt (shrK A) Number Tk, Agent B, Key K, X  parts (knows Spy evs);
        Notes Spy Number Tk, Key K  set evs;
        A  bad;  B  bad;  evs  bankerb_gets 
       Says B A (Crypt K (Number Ta))  set evs"
by (blast dest!: Kab_authentic
          intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE])

lemma B_authenticates_A:
     " Crypt K Agent A, Number Ta  parts (spies evs);
         Crypt (shrK B) Number Tk, Agent A, Key K   parts (spies evs);
        Key K  analz (spies evs);
         A  bad;  B  bad;  evs  bankerb_gets 
       Says A B Crypt (shrK B) Number Tk, Agent A, Key K,
                     Crypt K Agent A, Number Ta  set evs"
apply (blast dest!: ticket_authentic intro!: lemma_A)
done

lemma A_authenticates_B:
     " Crypt K (Number Ta)  parts (spies evs);
        Crypt (shrK A) Number Tk, Agent B, Key K, X  parts (spies evs);
        Key K  analz (spies evs);
        A  bad;  B  bad;  evs  bankerb_gets 
       Says B A (Crypt K (Number Ta))  set evs"
apply (blast dest!: Kab_authentic intro!: lemma_B)
done


subsection‹Temporal guarantees, relying on a temporal check that insures that
no oops event occurred. These are available in the sense of goal availability›


text‹Temporal treatment of confidentiality›

text‹Lemma: the session key sent in msg BK2 would be EXPIRED
    if the spy could see it!›
lemma lemma_conf_temporal [rule_format (no_asm)]:
     " A  bad;  B  bad;  evs  bankerb_gets 
   Says Server A
          (Crypt (shrK A) Number Tk, Agent B, Key K,
                            Crypt (shrK B) Number Tk, Agent A, Key K)
          set evs 
      Key K  analz (knows Spy evs)  expiredK Tk evs"
apply (erule bankerb_gets.induct)
apply (frule_tac [8] Says_Server_message_form)
apply (frule_tac [6] Gets_Server_message_form [THEN disjE])
apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)
txt‹Fake›
apply spy_analz
txt‹BK2›
apply (blast intro: parts_insertI less_SucI)
txt‹BK3›
apply (case_tac "Aa  bad")
 prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI)
txt‹Oops: PROOF FAILS if unsafe intro below›
apply (blast dest: unique_session_keys intro!: less_SucI)
done


text‹Confidentiality for the Server: Spy does not see the keys sent in msg BK2
as long as they have not expired.›
lemma Confidentiality_S_temporal:
     " Says Server A
          (Crypt K' Number T, Agent B, Key K, X)  set evs;
         ¬ expiredK T evs;
         A  bad;  B  bad;  evs  bankerb_gets
        Key K  analz (knows Spy evs)"
apply (frule Says_Server_message_form, assumption)
apply (blast intro: lemma_conf_temporal)
done

text‹Confidentiality for Alice›
lemma Confidentiality_A_temporal:
     " Crypt (shrK A) Number T, Agent B, Key K, X  parts (knows Spy evs);
         ¬ expiredK T evs;
         A  bad;  B  bad;  evs  bankerb_gets
        Key K  analz (knows Spy evs)"
by (blast dest!: Kab_authentic Confidentiality_S_temporal)

text‹Confidentiality for Bob›
lemma Confidentiality_B_temporal:
     " Crypt (shrK B) Number Tk, Agent A, Key K
           parts (knows Spy evs);
        ¬ expiredK Tk evs;
        A  bad;  B  bad;  evs  bankerb_gets
        Key K  analz (knows Spy evs)"
by (blast dest!: ticket_authentic Confidentiality_S_temporal)


text‹Temporal treatment of authentication›

text‹Authentication of A to B›
lemma B_authenticates_A_temporal:
     " Crypt K Agent A, Number Ta  parts (knows Spy evs);
         Crypt (shrK B) Number Tk, Agent A, Key K
          parts (knows Spy evs);
         ¬ expiredK Tk evs;
         A  bad;  B  bad;  evs  bankerb_gets 
       Says A B Crypt (shrK B) Number Tk, Agent A, Key K,
                     Crypt K Agent A, Number Ta  set evs"
by (blast dest!: ticket_authentic
          intro!: lemma_A
          elim!: Confidentiality_S_temporal [THEN [2] rev_notE])

text‹Authentication of B to A›
lemma A_authenticates_B_temporal:
     " Crypt K (Number Ta)  parts (knows Spy evs);
         Crypt (shrK A) Number Tk, Agent B, Key K, X
          parts (knows Spy evs);
         ¬ expiredK Tk evs;
         A  bad;  B  bad;  evs  bankerb_gets 
       Says B A (Crypt K (Number Ta))  set evs"
by (blast dest!: Kab_authentic
          intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE])


subsection‹Combined guarantees of key distribution and non-injective agreement on the session keys›

lemma B_authenticates_and_keydist_to_A:
     " Gets B Crypt (shrK B) Number Tk, Agent A, Key K,
                Crypt K Agent A, Number Ta  set evs;
        Key K  analz (spies evs);
        A  bad;  B  bad;  evs  bankerb_gets 
     Says A B Crypt (shrK B) Number Tk, Agent A, Key K,
                  Crypt K Agent A, Number Ta  set evs 
       Key K  analz (knows A evs)"
apply (blast dest: B_authenticates_A BK3_imp_Gets Gets_A_knows_K)
done

lemma A_authenticates_and_keydist_to_B:
     " Gets A (Crypt (shrK A) Number Tk, Agent B, Key K, Ticket)  set evs;
        Gets A (Crypt K (Number Ta))  set evs;
        Key K  analz (spies evs);
        A  bad;  B  bad;  evs  bankerb_gets 
     Says B A (Crypt K (Number Ta))  set evs
       Key K  analz (knows B evs)"
apply (blast dest: A_authenticates_B BK4_imp_Gets Gets_B_knows_K)
done





end