src/HOL/Auth/NS_Public.thy
changeset 11104 f2024fed9f0c
parent 5434 9b4bed3f394c
child 11230 756c5034f08b
equal deleted inserted replaced
11103:2a3cc8e1723a 11104:f2024fed9f0c
     5 
     5 
     6 Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol.
     6 Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol.
     7 Version incorporating Lowe's fix (inclusion of B's identity in round 2).
     7 Version incorporating Lowe's fix (inclusion of B's identity in round 2).
     8 *)
     8 *)
     9 
     9 
    10 NS_Public = Public + 
    10 theory NS_Public = Public:
    11 
    11 
    12 consts  ns_public  :: event list set
    12 consts  ns_public  :: "event list set"
    13 
    13 
    14 inductive ns_public
    14 inductive ns_public
    15   intrs 
    15   intros 
    16          (*Initial trace is empty*)
    16          (*Initial trace is empty*)
    17     Nil  "[]: ns_public"
    17    Nil:  "[] \<in> ns_public"
    18 
    18 
    19          (*The spy MAY say anything he CAN say.  We do not expect him to
    19          (*The spy MAY say anything he CAN say.  We do not expect him to
    20            invent new nonces here, but he can also use NS1.  Common to
    20            invent new nonces here, but he can also use NS1.  Common to
    21            all similar protocols.*)
    21            all similar protocols.*)
    22     Fake "[| evs: ns_public;  X: synth (analz (spies evs)) |]
    22    Fake: "\<lbrakk>evs \<in> ns_public;  X \<in> synth (analz (spies evs))\<rbrakk>
    23           ==> Says Spy B X  # evs : ns_public"
    23           \<Longrightarrow> Says Spy B X  # evs \<in> ns_public"
    24 
    24 
    25          (*Alice initiates a protocol run, sending a nonce to Bob*)
    25          (*Alice initiates a protocol run, sending a nonce to Bob*)
    26     NS1  "[| evs1: ns_public;  Nonce NA ~: used evs1 |]
    26    NS1:  "\<lbrakk>evs1 \<in> ns_public;  Nonce NA \<notin> used evs1\<rbrakk>
    27           ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|})
    27           \<Longrightarrow> Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
    28                  # evs1  :  ns_public"
    28                  # evs1  \<in>  ns_public"
    29 
    29 
    30          (*Bob responds to Alice's message with a further nonce*)
    30          (*Bob responds to Alice's message with a further nonce*)
    31     NS2  "[| evs2: ns_public;  Nonce NB ~: used evs2;
    31    NS2:  "\<lbrakk>evs2 \<in> ns_public;  Nonce NB \<notin> used evs2;
    32              Says A' B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs2 |]
    32            Says A' B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
    33           ==> Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})
    33           \<Longrightarrow> Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
    34                 # evs2  :  ns_public"
    34                 # evs2  \<in>  ns_public"
    35 
    35 
    36          (*Alice proves her existence by sending NB back to Bob.*)
    36          (*Alice proves her existence by sending NB back to Bob.*)
    37     NS3  "[| evs3: ns_public;
    37    NS3:  "\<lbrakk>evs3 \<in> ns_public;
    38              Says A  B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs3;
    38            Says A  B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
    39              Says B' A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})
    39            Says B' A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
    40                : set evs3 |]
    40               \<in> set evs3\<rbrakk>
    41           ==> Says A B (Crypt (pubK B) (Nonce NB)) # evs3 : ns_public"
    41           \<Longrightarrow> Says A B (Crypt (pubK B) (Nonce NB)) # evs3 \<in> ns_public"
    42 
    42 
    43   (**Oops message??**)
    43   (*No Oops message.  Should there be one?*)
       
    44 
       
    45 declare knows_Spy_partsEs [elim]
       
    46 declare analz_subset_parts [THEN subsetD, dest]
       
    47 declare Fake_parts_insert [THEN subsetD, dest]
       
    48 declare image_eq_UN [simp]  (*accelerates proofs involving nested images*)
       
    49 
       
    50 (*A "possibility property": there are traces that reach the end*)
       
    51 lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
       
    52 apply (intro exI bexI)
       
    53 apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2, 
       
    54                                    THEN ns_public.NS3])
       
    55 by possibility
       
    56 
       
    57 
       
    58 (**** Inductive proofs about ns_public ****)
       
    59 
       
    60 (** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY
       
    61     sends messages containing X! **)
       
    62 
       
    63 (*Spy never sees another agent's private key! (unless it's bad at start)*)
       
    64 lemma Spy_see_priK [simp]: 
       
    65       "evs \<in> ns_public \<Longrightarrow> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)"
       
    66 by (erule ns_public.induct, auto)
       
    67 
       
    68 lemma Spy_analz_priK [simp]: 
       
    69       "evs \<in> ns_public \<Longrightarrow> (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)"
       
    70 by auto
       
    71 
       
    72 
       
    73 (*** Authenticity properties obtained from NS2 ***)
       
    74 
       
    75 
       
    76 (*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce
       
    77   is secret.  (Honest users generate fresh nonces.)*)
       
    78 lemma no_nonce_NS1_NS2 [rule_format]: 
       
    79       "evs \<in> ns_public 
       
    80        \<Longrightarrow> Crypt (pubK C) \<lbrace>NA', Nonce NA, Agent D\<rbrace> \<in> parts (spies evs) \<longrightarrow>
       
    81            Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>  
       
    82            Nonce NA \<in> analz (spies evs)"
       
    83 apply (erule ns_public.induct, simp_all)
       
    84 apply (blast intro: analz_insertI)+
       
    85 done
       
    86 
       
    87 (*Unicity for NS1: nonce NA identifies agents A and B*)
       
    88 lemma unique_NA: 
       
    89      "\<lbrakk>Crypt(pubK B)  \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(spies evs);  
       
    90        Crypt(pubK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(spies evs);  
       
    91        Nonce NA \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>
       
    92       \<Longrightarrow> A=A' \<and> B=B'"
       
    93 apply (erule rev_mp, erule rev_mp, erule rev_mp)   
       
    94 apply (erule ns_public.induct, simp_all)
       
    95 (*Fake, NS1*)
       
    96 apply (blast intro: analz_insertI)+
       
    97 done
       
    98 
       
    99 
       
   100 (*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure
       
   101   The major premise "Says A B ..." makes it a dest-rule, so we use
       
   102   (erule rev_mp) rather than rule_format. *)
       
   103 theorem Spy_not_see_NA: 
       
   104       "\<lbrakk>Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
       
   105         A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
       
   106        \<Longrightarrow> Nonce NA \<notin> analz (spies evs)"
       
   107 apply (erule rev_mp)   
       
   108 apply (erule ns_public.induct, simp_all)
       
   109 apply spy_analz
       
   110 apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+
       
   111 done
       
   112 
       
   113 
       
   114 (*Authentication for A: if she receives message 2 and has used NA
       
   115   to start a run, then B has sent message 2.*)
       
   116 lemma A_trusts_NS2_lemma [rule_format]: 
       
   117    "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
       
   118     \<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs) \<longrightarrow>
       
   119 	Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
       
   120 	Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
       
   121 apply (erule ns_public.induct, simp_all)
       
   122 (*Fake, NS1*)
       
   123 apply (blast dest: Spy_not_see_NA)+
       
   124 done
       
   125 
       
   126 theorem A_trusts_NS2: 
       
   127      "\<lbrakk>Says A  B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;   
       
   128        Says B' A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
       
   129        A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
       
   130       \<Longrightarrow> Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
       
   131 by (blast intro: A_trusts_NS2_lemma)
       
   132 
       
   133 
       
   134 (*If the encrypted message appears then it originated with Alice in NS1*)
       
   135 lemma B_trusts_NS1 [rule_format]:
       
   136      "evs \<in> ns_public                                         
       
   137       \<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
       
   138 	  Nonce NA \<notin> analz (spies evs) \<longrightarrow>
       
   139 	  Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
       
   140 apply (erule ns_public.induct, simp_all)
       
   141 (*Fake*)
       
   142 apply (blast intro!: analz_insertI)
       
   143 done
       
   144 
       
   145 
       
   146 
       
   147 (*** Authenticity properties obtained from NS2 ***)
       
   148 
       
   149 (*Unicity for NS2: nonce NB identifies nonce NA and agents A, B 
       
   150   [unicity of B makes Lowe's fix work]
       
   151   [proof closely follows that for unique_NA] *)
       
   152 
       
   153 lemma unique_NB [dest]: 
       
   154      "\<lbrakk>Crypt(pubK A)  \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts(spies evs);
       
   155        Crypt(pubK A') \<lbrace>Nonce NA', Nonce NB, Agent B'\<rbrace> \<in> parts(spies evs);
       
   156        Nonce NB \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>
       
   157       \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B'"
       
   158 apply (erule rev_mp, erule rev_mp, erule rev_mp)   
       
   159 apply (erule ns_public.induct, simp_all)
       
   160 (*Fake, NS2*)
       
   161 apply (blast intro: analz_insertI)+
       
   162 done
       
   163 
       
   164 
       
   165 (*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)
       
   166 theorem Spy_not_see_NB [dest]:
       
   167      "\<lbrakk>Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
       
   168        A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
       
   169       \<Longrightarrow> Nonce NB \<notin> analz (spies evs)"
       
   170 apply (erule rev_mp)
       
   171 apply (erule ns_public.induct, simp_all)
       
   172 apply spy_analz
       
   173 apply (blast intro: no_nonce_NS1_NS2)+
       
   174 done
       
   175 
       
   176 
       
   177 (*Authentication for B: if he receives message 3 and has used NB
       
   178   in message 2, then A has sent message 3.*)
       
   179 lemma B_trusts_NS3_lemma [rule_format]:
       
   180      "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
       
   181       Crypt (pubK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
       
   182       Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
       
   183       Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
       
   184 by (erule ns_public.induct, auto)
       
   185 
       
   186 theorem B_trusts_NS3:
       
   187      "\<lbrakk>Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
       
   188        Says A' B (Crypt (pubK B) (Nonce NB)) \<in> set evs;             
       
   189        A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                    
       
   190       \<Longrightarrow> Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
       
   191 by (blast intro: B_trusts_NS3_lemma)
       
   192 
       
   193 (*** Overall guarantee for B ***)
       
   194 
       
   195 
       
   196 (*If NS3 has been sent and the nonce NB agrees with the nonce B joined with
       
   197   NA, then A initiated the run using NA.*)
       
   198 theorem B_trusts_protocol:
       
   199      "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
       
   200       Crypt (pubK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
       
   201       Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
       
   202       Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
       
   203 by (erule ns_public.induct, auto)
    44 
   204 
    45 end
   205 end