Theory NS_Public

(*  Title:      HOL/Auth/NS_Public.thy
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1996  University of Cambridge

Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol.
Version incorporating Lowe's fix (inclusion of B's identity in round 2).
*)

section‹Verifying the Needham-Schroeder-Lowe Public-Key Protocol›

theory NS_Public 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 (spies evsf))
           Says Spy B X  # evsf  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;
           Says A' B (Crypt (pubEK B) Nonce NA, Agent A)  set evs2
           Says B A (Crypt (pubEK A) Nonce NA, Nonce NB, Agent B)
                # 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;
           Says B' A (Crypt (pubEK A) Nonce NA, Nonce NB, Agent B)
               set evs3
           Says A B (Crypt (pubEK B) (Nonce NB)) # evs3  ns_public"

declare knows_Spy_partsEs [elim]
declare knows_Spy_partsEs [elim]
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.NS2, 
                                   THEN ns_public.NS3], possibility)
done

(** Theorems of the form X ∉ parts (spies 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 (spies evs)) = (A  bad)"
by (erule ns_public.induct, auto)

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

subsection‹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, Agent D  parts (spies evs) 
           Crypt (pubEK B) Nonce NA, Agent A  parts (spies evs)   
           Nonce NA  analz (spies 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(spies evs);  
       Crypt(pubEK B') Nonce NA, Agent A'  parts(spies evs);  
       Nonce NA  analz (spies 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 (spies 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, Agent B  parts (spies evs) 
        Says A B (Crypt(pubEK B) Nonce NA, Agent A)  set evs 
        Says B A (Crypt(pubEK A) Nonce NA, Nonce NB, Agent B)  set evs"
apply (erule ns_public.induct, simp_all)
(*Fake, NS1*)
apply (blast dest: Spy_not_see_NA)+
done

theorem A_trusts_NS2: 
     "Says A  B (Crypt(pubEK B) Nonce NA, Agent A)  set evs;   
       Says B' A (Crypt(pubEK A) Nonce NA, Nonce NB, Agent B)  set evs;
       A  bad;  B  bad;  evs  ns_public                     
       Says B A (Crypt(pubEK A) Nonce NA, Nonce NB, Agent B)  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 (spies evs) 
          Nonce NA  analz (spies 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


subsection‹Authenticity properties obtained from NS2›

(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B 
  [unicity of B makes Lowe's fix work]
  [proof closely follows that for unique_NA] *)

lemma unique_NB [dest]: 
     "Crypt(pubEK A)  Nonce NA, Nonce NB, Agent B  parts(spies evs);
       Crypt(pubEK A') Nonce NA', Nonce NB, Agent B'  parts(spies evs);
       Nonce NB  analz (spies evs); evs  ns_public
       A=A'  NA=NA'  B=B'"
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


(*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)
theorem Spy_not_see_NB [dest]:
     "Says B A (Crypt (pubEK A) Nonce NA, Nonce NB, Agent B)  set evs;
       A  bad;  B  bad;  evs  ns_public
       Nonce NB  analz (spies evs)"
apply (erule rev_mp)
apply (erule ns_public.induct, simp_all, spy_analz)
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.*)
lemma B_trusts_NS3_lemma [rule_format]:
     "A  bad;  B  bad;  evs  ns_public 
      Crypt (pubEK B) (Nonce NB)  parts (spies evs) 
      Says B A (Crypt (pubEK A) Nonce NA, Nonce NB, Agent B)  set evs 
      Says A B (Crypt (pubEK B) (Nonce NB))  set evs"
by (erule ns_public.induct, auto)

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

subsection‹Overall guarantee for B›

(*If NS3 has been sent and the nonce NB agrees with the nonce B joined with
  NA, then A initiated the run using NA.*)
theorem B_trusts_protocol:
     "A  bad;  B  bad;  evs  ns_public 
      Crypt (pubEK B) (Nonce NB)  parts (spies evs) 
      Says B A  (Crypt (pubEK A) Nonce NA, Nonce NB, Agent B)  set evs 
      Says A B (Crypt (pubEK B) Nonce NA, Agent A)  set evs"
by (erule ns_public.induct, auto)

end