Theory NS_Public_Bad

section‹The Needham-Schroeder Public-Key Protocol against Dolev-Yao --- with Gets event, hence with Reception rule›

theory NS_Public_Bad imports Public begin

inductive_set ns_public :: "event list set"
  where
         (*Initial trace is empty*)
   Nil:  "[]  ns_public"

         (*The spy MAY say anything he CAN say.  We do not expect him to
           invent new nonces here, but he can also use NS1.  Common to
           all similar protocols.*)
 | Fake: "evsf  ns_public;  X  synth (analz (knows Spy evsf))
           Says Spy B X  # evsf  ns_public"

         (*A message that is sent may be received*)
 | Reception: "evsr  ns_public; Says A B X  set evsr
                 Gets B X # evsr  ns_public"

         (*Alice initiates a protocol run, sending a nonce to Bob*)
 | NS1:  "evs1  ns_public;  Nonce NA  used evs1
           Says A B (Crypt (pubEK B) Nonce NA, Agent A)
                # evs1    ns_public"

         (*Bob responds to Alice's message with a further nonce*)
 | NS2:  "evs2  ns_public;  Nonce NB  used evs2;
           Gets B (Crypt (pubEK B) Nonce NA, Agent A)  set evs2
           Says B A (Crypt (pubEK A) Nonce NA, Nonce NB)
                # evs2    ns_public"

         (*Alice proves her existence by sending NB back to Bob.*)
 | NS3:  "evs3  ns_public;
           Says A  B (Crypt (pubEK B) Nonce NA, Agent A)  set evs3;
           Gets A (Crypt (pubEK A) Nonce NA, Nonce NB)  set evs3
           Says A B (Crypt (pubEK B) (Nonce NB)) # evs3  ns_public"

declare knows_Spy_partsEs [elim] thm knows_Spy_partsEs
declare analz_into_parts [dest]
declare Fake_parts_insert_in_Un [dest]

(*A "possibility property": there are traces that reach the end*)
lemma "NB. evs  ns_public. Says A B (Crypt (pubEK B) (Nonce NB))  set evs"
apply (intro exI bexI)
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.Reception, 
                                   THEN ns_public.NS2, THEN ns_public.Reception, 
                                   THEN ns_public.NS3])
by possibility


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

lemma Gets_imp_knows_Spy: 
     " Gets B X  set evs; evs  ns_public    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  ns_public    X  parts (knows Spy evs)"
apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj])
done

(**** Inductive proofs about ns_public ****)

(** Theorems of the form X ∉ parts (knows Spy evs) imply that NOBODY
    sends messages containing X! **)

(*Spy never sees another agent's private key! (unless it's bad at start)*)
lemma Spy_see_priEK [simp]: 
      "evs  ns_public  (Key (priEK A)  parts (knows Spy evs)) = (A  bad)"
by (erule ns_public.induct, auto)

lemma Spy_analz_priEK [simp]: 
      "evs  ns_public  (Key (priEK A)  analz (knows Spy evs)) = (A  bad)"
by auto


(*** Authenticity properties obtained from NS2 ***)

(*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce
  is secret.  (Honest users generate fresh nonces.)*)
lemma no_nonce_NS1_NS2 [rule_format]: 
      "evs  ns_public 
        Crypt (pubEK C) NA', Nonce NA  parts (knows Spy evs) 
           Crypt (pubEK B) Nonce NA, Agent A  parts (knows Spy evs)   
           Nonce NA  analz (knows Spy evs)"
apply (erule ns_public.induct, simp_all)
apply (blast intro: analz_insertI)+
done


(*Unicity for NS1: nonce NA identifies agents A and B*)
lemma unique_NA: 
     "Crypt(pubEK B)  Nonce NA, Agent A   parts(knows Spy evs);  
       Crypt(pubEK B') Nonce NA, Agent A'  parts(knows Spy evs);  
       Nonce NA  analz (knows Spy evs); evs  ns_public
       A=A'  B=B'"
apply (erule rev_mp, erule rev_mp, erule rev_mp)   
apply (erule ns_public.induct, simp_all)
(*Fake, NS1*)
apply (blast intro!: analz_insertI)+
done


(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure
  The major premise "Says A B ..." makes it a dest-rule, so we use
  (erule rev_mp) rather than rule_format. *)
theorem Spy_not_see_NA: 
      "Says A B (Crypt(pubEK B) Nonce NA, Agent A)  set evs;
        A  bad;  B  bad;  evs  ns_public                     
        Nonce NA  analz (knows Spy evs)"
apply (erule rev_mp)   
apply (erule ns_public.induct, simp_all, spy_analz)
apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+
done


(*Authentication for A: if she receives message 2 and has used NA
  to start a run, then B has sent message 2.*)
lemma A_trusts_NS2_lemma [rule_format]: 
   "A  bad;  B  bad;  evs  ns_public                     
     Crypt (pubEK A) Nonce NA, Nonce NB  parts (knows Spy evs) 
        Says A B (Crypt(pubEK B) Nonce NA, Agent A)  set evs 
        Says B A (Crypt(pubEK A) Nonce NA, Nonce NB)  set evs"
apply (erule ns_public.induct)
apply (auto dest: Spy_not_see_NA unique_NA)
done

theorem A_trusts_NS2: 
     "Says A  B (Crypt(pubEK B) Nonce NA, Agent A)  set evs;   
       Gets A (Crypt(pubEK A) Nonce NA, Nonce NB)  set evs;
       A  bad;  B  bad;  evs  ns_public                     
       Says B A (Crypt(pubEK A) Nonce NA, Nonce NB)  set evs"
by (blast intro: A_trusts_NS2_lemma)


(*If the encrypted message appears then it originated with Alice in NS1*)
lemma B_trusts_NS1 [rule_format]:
     "evs  ns_public                                         
       Crypt (pubEK B) Nonce NA, Agent A  parts (knows Spy evs) 
          Nonce NA  analz (knows Spy evs) 
          Says A B (Crypt (pubEK B) Nonce NA, Agent A)  set evs"
apply (erule ns_public.induct, simp_all)
(*Fake*)
apply (blast intro!: analz_insertI)
done



(*** Authenticity properties obtained from NS2 ***)

(*Unicity for NS2: nonce NB identifies nonce NA and agent A
  [proof closely follows that for unique_NA] *)
lemma unique_NB [dest]: 
     "Crypt(pubEK A)  Nonce NA, Nonce NB  parts(knows Spy evs);
       Crypt(pubEK A') Nonce NA', Nonce NB  parts(knows Spy evs);
       Nonce NB  analz (knows Spy evs); evs  ns_public
      A=A'  NA=NA'"
apply (erule rev_mp, erule rev_mp, erule rev_mp)   
apply (erule ns_public.induct, simp_all)
(*Fake, NS2*)
apply (blast intro!: analz_insertI)+
done


(*NB remains secret PROVIDED Alice never responds with round 3*)
theorem Spy_not_see_NB [dest]:
     "Says B A (Crypt (pubEK A) Nonce NA, Nonce NB)  set evs;   
       C. Says A C (Crypt (pubEK C) (Nonce NB))  set evs;       
       A  bad;  B  bad;  evs  ns_public                      
      Nonce NB  analz (knows Spy evs)"
apply (erule rev_mp, erule rev_mp)
apply (erule ns_public.induct, simp_all, spy_analz)
apply (simp_all add: all_conj_distrib) (*speeds up the next step*)
apply (blast intro: no_nonce_NS1_NS2)+
done


(*Authentication for B: if he receives message 3 and has used NB
  in message 2, then A has sent message 3--to somebody....*)

lemma B_trusts_NS3_lemma [rule_format]:
     "A  bad;  B  bad;  evs  ns_public                    
       Crypt (pubEK B) (Nonce NB)  parts (knows Spy evs) 
          Says B A  (Crypt (pubEK A) Nonce NA, Nonce NB)  set evs 
          (C. Says A C (Crypt (pubEK C) (Nonce NB))  set evs)"
apply (erule ns_public.induct, auto)
by (blast intro: no_nonce_NS1_NS2)+

theorem B_trusts_NS3:
     "Says B A  (Crypt (pubEK A) Nonce NA, Nonce NB)  set evs;
       Gets B (Crypt (pubEK B) (Nonce NB))  set evs;             
       A  bad;  B  bad;  evs  ns_public                    
       C. Says A C (Crypt (pubEK C) (Nonce NB))  set evs"
by (blast intro: B_trusts_NS3_lemma)


(*Can we strengthen the secrecy theorem Spy_not_see_NB?  NO*)
lemma "A  bad;  B  bad;  evs  ns_public            
        Says B A (Crypt (pubEK A) Nonce NA, Nonce NB)  set evs  
            Nonce NB  analz (knows Spy evs)"
apply (erule ns_public.induct, simp_all, spy_analz)
(*NS1: by freshness*)
apply blast
(*NS2: by freshness and unicity of NB*)
apply (blast intro: no_nonce_NS1_NS2)
(*NS3: unicity of NB identifies A and NA, but not B*)
apply clarify
apply (frule_tac A' = A in 
       Says_imp_knows_Spy [THEN parts.Inj, THEN unique_NB], auto)
apply (rename_tac evs3 B' C)
txt‹This is the attack!
@{subgoals[display,indent=0,margin=65]}
oops

end