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