Theory Recur
section‹The Otway-Bull Recursive Authentication Protocol›
theory Recur imports Public begin
text‹End marker for message bundles›
abbreviation
  END :: "msg" where
  "END == Number 0"
inductive_set 
  respond :: "event list ⇒ (msg*msg*key)set"
  for evs :: "event list"
  where
   One:  "Key KAB ∉ used evs
          ⟹ (Hash[Key(shrK A)] ⦃Agent A, Agent B, Nonce NA, END⦄,
               ⦃Crypt (shrK A) ⦃Key KAB, Agent B, Nonce NA⦄, END⦄,
               KAB)   ∈ respond evs"
    
 | Cons: "⟦(PA, RA, KAB) ∈ respond evs;
             Key KBC ∉ used evs;  Key KBC ∉ parts {RA};
             PA = Hash[Key(shrK A)] ⦃Agent A, Agent B, Nonce NA, P⦄⟧
          ⟹ (Hash[Key(shrK B)] ⦃Agent B, Agent C, Nonce NB, PA⦄,
               ⦃Crypt (shrK B) ⦃Key KBC, Agent C, Nonce NB⦄,
                 Crypt (shrK B) ⦃Key KAB, Agent A, Nonce NB⦄,
                 RA⦄,
               KBC)
              ∈ respond evs"
inductive_set
  responses :: "event list => msg set"
  for evs :: "event list"
  where
    
   Nil:  "END ∈ responses evs"
 | Cons: "⟦RA ∈ responses evs;  Key KAB ∉ used evs⟧
          ⟹ ⦃Crypt (shrK B) ⦃Key KAB, Agent A, Nonce NB⦄,
                RA⦄  ∈ responses evs"
inductive_set recur :: "event list set"
  where
         
   Nil:  "[] ∈ recur"
         
 | Fake: "⟦evsf ∈ recur;  X ∈ synth (analz (knows Spy evsf))⟧
          ⟹ Says Spy B X  # evsf ∈ recur"
         
 | RA1:  "⟦evs1 ∈ recur;  Nonce NA ∉ used evs1⟧
          ⟹ Says A B (Hash[Key(shrK A)] ⦃Agent A, Agent B, Nonce NA, END⦄)
              # evs1 ∈ recur"
         
 | RA2:  "⟦evs2 ∈ recur;  Nonce NB ∉ used evs2;
             Says A' B PA ∈ set evs2⟧
          ⟹ Says B C (Hash[Key(shrK B)] ⦃Agent B, Agent C, Nonce NB, PA⦄)
              # evs2 ∈ recur"
         
 | RA3:  "⟦evs3 ∈ recur;  Says B' Server PB ∈ set evs3;
             (PB,RB,K) ∈ respond evs3⟧
          ⟹ Says Server B RB # evs3 ∈ recur"
         
 | RA4:  "⟦evs4 ∈ recur;
             Says B  C ⦃XH, Agent B, Agent C, Nonce NB,
                         XA, Agent A, Agent B, Nonce NA, P⦄ ∈ set evs4;
             Says C' B ⦃Crypt (shrK B) ⦃Key KBC, Agent C, Nonce NB⦄,
                         Crypt (shrK B) ⦃Key KAB, Agent A, Nonce NB⦄,
                         RA⦄ ∈ set evs4⟧
          ⟹ Says B A RA # evs4 ∈ recur"
   
declare Says_imp_knows_Spy [THEN analz.Inj, dest]
declare parts.Body  [dest]
declare analz_into_parts [dest]
declare Fake_parts_insert_in_Un  [dest]
text‹Simplest case: Alice goes directly to the server›
lemma "Key K ∉ used [] 
       ⟹ ∃NA. ∃evs ∈ recur.
              Says Server A ⦃Crypt (shrK A) ⦃Key K, Agent Server, Nonce NA⦄,
                    END⦄  ∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] recur.Nil [THEN recur.RA1, 
                             THEN recur.RA3 [OF _ _ respond.One]])
apply (possibility, simp add: used_Cons) 
done
text‹Case two: Alice, Bob and the server›
lemma "⟦Key K ∉ used []; Key K' ∉ used []; K ≠ K';
          Nonce NA ∉ used []; Nonce NB ∉ used []; NA < NB⟧
       ⟹ ∃NA. ∃evs ∈ recur.
        Says B A ⦃Crypt (shrK A) ⦃Key K, Agent B, Nonce NA⦄,
                   END⦄  ∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] 
          recur.Nil
           [THEN recur.RA1 [of _ NA], 
            THEN recur.RA2 [of _ NB],
            THEN recur.RA3 [OF _ _ respond.One 
                                     [THEN respond.Cons [of _ _ K _ K']]],
            THEN recur.RA4], possibility)
apply (auto simp add: used_Cons)
done
lemma "⟦Key K ∉ used []; Key K' ∉ used [];  
          Key K'' ∉ used []; K ≠ K'; K' ≠ K''; K ≠ K'';
          Nonce NA ∉ used []; Nonce NB ∉ used []; Nonce NC ∉ used []; 
          NA < NB; NB < NC⟧
       ⟹ ∃K. ∃NA. ∃evs ∈ recur.
             Says B A ⦃Crypt (shrK A) ⦃Key K, Agent B, Nonce NA⦄,
                        END⦄  ∈ set evs"
apply (intro exI bexI)
apply (rule_tac [2] 
          recur.Nil [THEN recur.RA1, 
                     THEN recur.RA2, THEN recur.RA2,
                     THEN recur.RA3 
                          [OF _ _ respond.One 
                                  [THEN respond.Cons, THEN respond.Cons]],
                     THEN recur.RA4, THEN recur.RA4])
apply basic_possibility
apply (tactic "DEPTH_SOLVE (swap_res_tac \<^context> [refl, conjI, disjCI] 1)")
done
lemma respond_imp_not_used: "(PA,RB,KAB) ∈ respond evs ⟹ Key KAB ∉ used evs"
by (erule respond.induct, simp_all)
lemma Key_in_parts_respond [rule_format]:
   "⟦Key K ∈ parts {RB};  (PB,RB,K') ∈ respond evs⟧ ⟹ Key K ∉ used evs"
apply (erule rev_mp, erule respond.induct)
apply (auto dest: Key_not_used respond_imp_not_used)
done
text‹Simple inductive reasoning about responses›
lemma respond_imp_responses:
     "(PA,RB,KAB) ∈ respond evs ⟹ RB ∈ responses evs"
apply (erule respond.induct)
apply (blast intro!: respond_imp_not_used responses.intros)+
done
lemmas RA2_analz_spies = Says_imp_spies [THEN analz.Inj]
lemma RA4_analz_spies:
     "Says C' B ⦃Crypt K X, X', RA⦄ ∈ set evs ⟹ RA ∈ analz (spies evs)"
by blast
lemmas RA2_parts_spies =  RA2_analz_spies [THEN analz_into_parts]
lemmas RA4_parts_spies =  RA4_analz_spies [THEN analz_into_parts]
lemma Spy_see_shrK [simp]:
     "evs ∈ recur ⟹ (Key (shrK A) ∈ parts (spies evs)) = (A ∈ bad)"
apply (erule recur.induct, auto)
txt‹RA3.  It's ugly to call auto twice, but it seems necessary.›
apply (auto dest: Key_in_parts_respond simp add: parts_insert_spies)
done
lemma Spy_analz_shrK [simp]:
     "evs ∈ recur ⟹ (Key (shrK A) ∈ analz (spies evs)) = (A ∈ bad)"
by auto
lemma Spy_see_shrK_D [dest!]:
     "⟦Key (shrK A) ∈ parts (knows Spy evs);  evs ∈ recur⟧ ⟹ A ∈ bad"
by (blast dest: Spy_see_shrK)
lemma resp_analz_image_freshK_lemma:
     "⟦RB ∈ responses evs;
         ∀K KK. KK ⊆ - (range shrK) ⟶
                   (Key K ∈ analz (Key`KK ∪ H)) =
                   (K ∈ KK | Key K ∈ analz H)⟧
     ⟹ ∀K KK. KK ⊆ - (range shrK) ⟶
                   (Key K ∈ analz (insert RB (Key`KK ∪ H))) =
                   (K ∈ KK | Key K ∈ analz (insert RB H))"
apply (erule responses.induct)
apply (simp_all del: image_insert
                add: analz_image_freshK_simps, auto)
done 
text‹Version for the protocol.  Proof is easy, thanks to the lemma.›
lemma raw_analz_image_freshK:
 "evs ∈ recur ⟹
   ∀K KK. KK ⊆ - (range shrK) ⟶
          (Key K ∈ analz (Key`KK ∪ (spies evs))) =
          (K ∈ KK | Key K ∈ analz (spies evs))"
apply (erule recur.induct)
apply (drule_tac [4] RA2_analz_spies,
       drule_tac [5] respond_imp_responses,
       drule_tac [6] RA4_analz_spies, analz_freshK, spy_analz)
txt‹RA3›
apply (simp_all add: resp_analz_image_freshK_lemma)
done
lemmas resp_analz_image_freshK =  
       resp_analz_image_freshK_lemma [OF _ raw_analz_image_freshK]
lemma analz_insert_freshK:
     "⟦evs ∈ recur;  KAB ∉ range shrK⟧
      ⟹ (Key K ∈ analz (insert (Key KAB) (spies evs))) =
          (K = KAB | Key K ∈ analz (spies evs))"
by (simp del: image_insert
         add: analz_image_freshK_simps raw_analz_image_freshK)
text‹Everything that's hashed is already in past traffic.›
lemma Hash_imp_body:
     "⟦Hash ⦃Key(shrK A), X⦄ ∈ parts (spies evs);
         evs ∈ recur;  A ∉ bad⟧ ⟹ X ∈ parts (spies evs)"
apply (erule rev_mp)
apply (erule recur.induct,
       drule_tac [6] RA4_parts_spies,
       drule_tac [5] respond_imp_responses,
       drule_tac [4] RA2_parts_spies)
txt‹RA3 requires a further induction›
apply (erule_tac [5] responses.induct, simp_all)
txt‹Fake›
apply (blast intro: parts_insertI)
done
lemma unique_NA:
  "⟦Hash ⦃Key(shrK A), Agent A, B, NA, P⦄ ∈ parts (spies evs);
      Hash ⦃Key(shrK A), Agent A, B',NA, P'⦄ ∈ parts (spies evs);
      evs ∈ recur;  A ∉ bad⟧
    ⟹ B=B' ∧ P=P'"
apply (erule rev_mp, erule rev_mp)
apply (erule recur.induct,
       drule_tac [5] respond_imp_responses)
apply (force, simp_all)
txt‹Fake›
apply blast
apply (erule_tac [3] responses.induct)
txt‹RA1,2: creation of new Nonce›
apply simp_all
apply (blast dest!: Hash_imp_body)+
done
lemma shrK_in_analz_respond [simp]:
     "⟦RB ∈ responses evs;  evs ∈ recur⟧
  ⟹ (Key (shrK B) ∈ analz (insert RB (spies evs))) = (B∈bad)"
apply (erule responses.induct)
apply (simp_all del: image_insert
                add: analz_image_freshK_simps resp_analz_image_freshK, auto) 
done
lemma resp_analz_insert_lemma:
     "⟦Key K ∈ analz (insert RB H);
         ∀K KK. KK ⊆ - (range shrK) ⟶
                   (Key K ∈ analz (Key`KK ∪ H)) =
                   (K ∈ KK | Key K ∈ analz H);
         RB ∈ responses evs⟧
     ⟹ (Key K ∈ parts{RB} | Key K ∈ analz H)"
apply (erule rev_mp, erule responses.induct)
apply (simp_all del: image_insert parts_image
             add: analz_image_freshK_simps resp_analz_image_freshK_lemma)
txt‹Simplification using two distinct treatments of "image"›
apply (simp add: parts_insert2, blast)
done
lemmas resp_analz_insert =
       resp_analz_insert_lemma [OF _ raw_analz_image_freshK]
text‹The last key returned by respond indeed appears in a certificate›
lemma respond_certificate:
     "(Hash[Key(shrK A)] ⦃Agent A, B, NA, P⦄, RA, K) ∈ respond evs
      ⟹ Crypt (shrK A) ⦃Key K, B, NA⦄ ∈ parts {RA}"
apply (ind_cases "(Hash[Key (shrK A)] ⦃Agent A, B, NA, P⦄, RA, K) ∈ respond evs")
apply simp_all
done
lemma unique_lemma [rule_format]:
     "(PB,RB,KXY) ∈ respond evs ⟹
      ∀A B N. Crypt (shrK A) ⦃Key K, Agent B, N⦄ ∈ parts {RB} ⟶
      (∀A' B' N'. Crypt (shrK A') ⦃Key K, Agent B', N'⦄ ∈ parts {RB} ⟶
      (A'=A ∧ B'=B) | (A'=B ∧ B'=A))"
apply (erule respond.induct)
apply (simp_all add: all_conj_distrib)
apply (blast dest: respond_certificate)
done
lemma unique_session_keys:
     "⟦Crypt (shrK A) ⦃Key K, Agent B, N⦄ ∈ parts {RB};
         Crypt (shrK A') ⦃Key K, Agent B', N'⦄ ∈ parts {RB};
         (PB,RB,KXY) ∈ respond evs⟧
      ⟹ (A'=A ∧ B'=B) | (A'=B ∧ B'=A)"
by (rule unique_lemma, auto)
lemma respond_Spy_not_see_session_key [rule_format]:
     "⟦(PB,RB,KAB) ∈ respond evs;  evs ∈ recur⟧
      ⟹ ∀A A' N. A ∉ bad ∧ A' ∉ bad ⟶
          Crypt (shrK A) ⦃Key K, Agent A', N⦄ ∈ parts{RB} ⟶
          Key K ∉ analz (insert RB (spies evs))"
apply (erule respond.induct)
apply (frule_tac [2] respond_imp_responses)
apply (frule_tac [2] respond_imp_not_used)
apply (simp_all del: image_insert parts_image
                add: analz_image_freshK_simps split_ifs shrK_in_analz_respond
                     resp_analz_image_freshK parts_insert2)
txt‹Base case of respond›
apply blast
txt‹Inductive step of respond›
apply (intro allI conjI impI, simp_all)
txt‹by unicity, either \<^term>‹B=Aa› or \<^term>‹B=A'›, a contradiction
     if \<^term>‹B ∈ bad››   
apply (blast dest: unique_session_keys respond_certificate)
apply (blast dest!: respond_certificate)
apply (blast dest!: resp_analz_insert)
done
lemma Spy_not_see_session_key:
     "⟦Crypt (shrK A) ⦃Key K, Agent A', N⦄ ∈ parts (spies evs);
         A ∉ bad;  A' ∉ bad;  evs ∈ recur⟧
      ⟹ Key K ∉ analz (spies evs)"
apply (erule rev_mp)
apply (erule recur.induct)
apply (drule_tac [4] RA2_analz_spies,
       frule_tac [5] respond_imp_responses,
       drule_tac [6] RA4_analz_spies,
       simp_all add: split_ifs analz_insert_eq analz_insert_freshK)
txt‹Fake›
apply spy_analz
txt‹RA2›
apply blast 
txt‹RA3›
apply (simp add: parts_insert_spies)
apply (metis Key_in_parts_respond parts.Body parts.Fst resp_analz_insert 
             respond_Spy_not_see_session_key usedI)
txt‹RA4›
apply blast 
done
text‹The response never contains Hashes›
lemma Hash_in_parts_respond:
     "⟦Hash ⦃Key (shrK B), M⦄ ∈ parts (insert RB H);
         (PB,RB,K) ∈ respond evs⟧
      ⟹ Hash ⦃Key (shrK B), M⦄ ∈ parts H"
apply (erule rev_mp)
apply (erule respond_imp_responses [THEN responses.induct], auto)
done
text‹Only RA1 or RA2 can have caused such a part of a message to appear.
  This result is of no use to B, who cannot verify the Hash.  Moreover,
  it can say nothing about how recent A's message is.  It might later be
  used to prove B's presence to A at the run's conclusion.›
lemma Hash_auth_sender [rule_format]:
     "⟦Hash ⦃Key(shrK A), Agent A, Agent B, NA, P⦄ ∈ parts(spies evs);
         A ∉ bad;  evs ∈ recur⟧
      ⟹ Says A B (Hash[Key(shrK A)] ⦃Agent A, Agent B, NA, P⦄) ∈ set evs"
unfolding HPair_def
apply (erule rev_mp)
apply (erule recur.induct,
       drule_tac [6] RA4_parts_spies,
       drule_tac [4] RA2_parts_spies,
       simp_all)
txt‹Fake, RA3›
apply (blast dest: Hash_in_parts_respond)+
done
text‹Certificates can only originate with the Server.›
lemma Cert_imp_Server_msg:
     "⟦Crypt (shrK A) Y ∈ parts (spies evs);
         A ∉ bad;  evs ∈ recur⟧
      ⟹ ∃C RC. Says Server C RC ∈ set evs  ∧
                   Crypt (shrK A) Y ∈ parts {RC}"
apply (erule rev_mp, erule recur.induct, simp_all)
txt‹Fake›
apply blast
txt‹RA1›
apply blast
txt‹RA2: it cannot be a new Nonce, contradiction.›
apply blast
txt‹RA3.  Pity that the proof is so brittle: this step requires the rewriting,
       which however would break all other steps.›
apply (simp add: parts_insert_spies, blast)
txt‹RA4›
apply blast
done
end