src/HOL/Auth/NS_Public.ML
changeset 3683 aafe719dff14
parent 3675 70dd312b70b2
child 3709 c13c2137e9ee
equal deleted inserted replaced
3682:597efdb7decb 3683:aafe719dff14
    10 open NS_Public;
    10 open NS_Public;
    11 
    11 
    12 proof_timing:=true;
    12 proof_timing:=true;
    13 HOL_quantifiers := false;
    13 HOL_quantifiers := false;
    14 
    14 
    15 AddIffs [Spy_in_lost];
    15 AddIffs [Spy_in_bad];
    16 
    16 
    17 (*A "possibility property": there are traces that reach the end*)
    17 (*A "possibility property": there are traces that reach the end*)
    18 goal thy 
    18 goal thy 
    19  "!!A B. A ~= B ==> EX NB. EX evs: ns_public.               \
    19  "!!A B. A ~= B ==> EX NB. EX evs: ns_public.               \
    20 \                     Says A B (Crypt (pubK B) (Nonce NB)) : set evs";
    20 \                     Says A B (Crypt (pubK B) (Nonce NB)) : set evs";
    34 Addsimps [not_Says_to_self];
    34 Addsimps [not_Says_to_self];
    35 AddSEs   [not_Says_to_self RSN (2, rev_notE)];
    35 AddSEs   [not_Says_to_self RSN (2, rev_notE)];
    36 
    36 
    37 
    37 
    38 (*Induction for regularity theorems.  If induction formula has the form
    38 (*Induction for regularity theorems.  If induction formula has the form
    39    X ~: analz (sees Spy evs) --> ... then it shortens the proof by discarding
    39    X ~: analz (spies evs) --> ... then it shortens the proof by discarding
    40    needless information about analz (insert X (sees Spy evs))  *)
    40    needless information about analz (insert X (spies evs))  *)
    41 fun parts_induct_tac i = 
    41 fun parts_induct_tac i = 
    42     etac ns_public.induct i
    42     etac ns_public.induct i
    43     THEN 
    43     THEN 
    44     REPEAT (FIRSTGOAL analz_mono_contra_tac)
    44     REPEAT (FIRSTGOAL analz_mono_contra_tac)
    45     THEN 
    45     THEN 
    46     prove_simple_subgoals_tac i;
    46     prove_simple_subgoals_tac i;
    47 
    47 
    48 
    48 
    49 (** Theorems of the form X ~: parts (sees Spy evs) imply that NOBODY
    49 (** Theorems of the form X ~: parts (spies evs) imply that NOBODY
    50     sends messages containing X! **)
    50     sends messages containing X! **)
    51 
    51 
    52 (*Spy never sees another agent's private key! (unless it's lost at start)*)
    52 (*Spy never sees another agent's private key! (unless it's bad at start)*)
    53 goal thy 
    53 goal thy 
    54  "!!A. evs: ns_public ==> (Key (priK A) : parts (sees Spy evs)) = (A : lost)";
    54  "!!A. evs: ns_public ==> (Key (priK A) : parts (spies evs)) = (A : bad)";
    55 by (parts_induct_tac 1);
    55 by (parts_induct_tac 1);
    56 by (Fake_parts_insert_tac 1);
    56 by (Fake_parts_insert_tac 1);
    57 qed "Spy_see_priK";
    57 qed "Spy_see_priK";
    58 Addsimps [Spy_see_priK];
    58 Addsimps [Spy_see_priK];
    59 
    59 
    60 goal thy 
    60 goal thy 
    61  "!!A. evs: ns_public ==> (Key (priK A) : analz (sees Spy evs)) = (A : lost)";
    61  "!!A. evs: ns_public ==> (Key (priK A) : analz (spies evs)) = (A : bad)";
    62 by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset));
    62 by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset));
    63 qed "Spy_analz_priK";
    63 qed "Spy_analz_priK";
    64 Addsimps [Spy_analz_priK];
    64 Addsimps [Spy_analz_priK];
    65 
    65 
    66 goal thy  "!!A. [| Key (priK A) : parts (sees Spy evs);       \
    66 goal thy  "!!A. [| Key (priK A) : parts (spies evs);       \
    67 \                  evs : ns_public |] ==> A:lost";
    67 \                  evs : ns_public |] ==> A:bad";
    68 by (blast_tac (!claset addDs [Spy_see_priK]) 1);
    68 by (blast_tac (!claset addDs [Spy_see_priK]) 1);
    69 qed "Spy_see_priK_D";
    69 qed "Spy_see_priK_D";
    70 
    70 
    71 bind_thm ("Spy_analz_priK_D", analz_subset_parts RS subsetD RS Spy_see_priK_D);
    71 bind_thm ("Spy_analz_priK_D", analz_subset_parts RS subsetD RS Spy_see_priK_D);
    72 AddSDs [Spy_see_priK_D, Spy_analz_priK_D];
    72 AddSDs [Spy_see_priK_D, Spy_analz_priK_D];
    75 (**** Authenticity properties obtained from NS2 ****)
    75 (**** Authenticity properties obtained from NS2 ****)
    76 
    76 
    77 (*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce
    77 (*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce
    78   is secret.  (Honest users generate fresh nonces.)*)
    78   is secret.  (Honest users generate fresh nonces.)*)
    79 goal thy 
    79 goal thy 
    80  "!!evs. [| Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees Spy evs); \
    80  "!!evs. [| Crypt (pubK B) {|Nonce NA, Agent A|} : parts (spies evs); \
    81 \           Nonce NA ~: analz (sees Spy evs);       \
    81 \           Nonce NA ~: analz (spies evs);       \
    82 \           evs : ns_public |]                      \
    82 \           evs : ns_public |]                      \
    83 \ ==> Crypt (pubK C) {|NA', Nonce NA, Agent D|} ~: parts (sees Spy evs)";
    83 \ ==> Crypt (pubK C) {|NA', Nonce NA, Agent D|} ~: parts (spies evs)";
    84 by (etac rev_mp 1);
    84 by (etac rev_mp 1);
    85 by (etac rev_mp 1);
    85 by (etac rev_mp 1);
    86 by (parts_induct_tac 1);
    86 by (parts_induct_tac 1);
    87 (*NS3*)
    87 (*NS3*)
    88 by (blast_tac (!claset addSEs partsEs) 3);
    88 by (blast_tac (!claset addSEs partsEs) 3);
    92 qed "no_nonce_NS1_NS2";
    92 qed "no_nonce_NS1_NS2";
    93 
    93 
    94 
    94 
    95 (*Unicity for NS1: nonce NA identifies agents A and B*)
    95 (*Unicity for NS1: nonce NA identifies agents A and B*)
    96 goal thy 
    96 goal thy 
    97  "!!evs. [| Nonce NA ~: analz (sees Spy evs);  evs : ns_public |]      \
    97  "!!evs. [| Nonce NA ~: analz (spies evs);  evs : ns_public |]      \
    98 \ ==> EX A' B'. ALL A B.                                               \
    98 \ ==> EX A' B'. ALL A B.                                               \
    99 \      Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees Spy evs) --> \
    99 \      Crypt (pubK B) {|Nonce NA, Agent A|} : parts (spies evs) --> \
   100 \      A=A' & B=B'";
   100 \      A=A' & B=B'";
   101 by (etac rev_mp 1);
   101 by (etac rev_mp 1);
   102 by (parts_induct_tac 1);
   102 by (parts_induct_tac 1);
   103 by (ALLGOALS
   103 by (ALLGOALS
   104     (asm_simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_sees])));
   104     (asm_simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_spies])));
   105 (*NS1*)
   105 (*NS1*)
   106 by (expand_case_tac "NA = ?y" 2 THEN blast_tac (!claset addSEs partsEs) 2);
   106 by (expand_case_tac "NA = ?y" 2 THEN blast_tac (!claset addSEs partsEs) 2);
   107 (*Fake*)
   107 (*Fake*)
   108 by (step_tac (!claset addSIs [analz_insertI]) 1);
   108 by (step_tac (!claset addSIs [analz_insertI]) 1);
   109 by (ex_strip_tac 1);
   109 by (ex_strip_tac 1);
   110 by (Fake_parts_insert_tac 1);
   110 by (Fake_parts_insert_tac 1);
   111 val lemma = result();
   111 val lemma = result();
   112 
   112 
   113 goal thy 
   113 goal thy 
   114  "!!evs. [| Crypt(pubK B)  {|Nonce NA, Agent A|}  : parts(sees Spy evs); \
   114  "!!evs. [| Crypt(pubK B)  {|Nonce NA, Agent A|}  : parts(spies evs); \
   115 \           Crypt(pubK B') {|Nonce NA, Agent A'|} : parts(sees Spy evs); \
   115 \           Crypt(pubK B') {|Nonce NA, Agent A'|} : parts(spies evs); \
   116 \           Nonce NA ~: analz (sees Spy evs);                            \
   116 \           Nonce NA ~: analz (spies evs);                            \
   117 \           evs : ns_public |]                                                \
   117 \           evs : ns_public |]                                        \
   118 \        ==> A=A' & B=B'";
   118 \        ==> A=A' & B=B'";
   119 by (prove_unique_tac lemma 1);
   119 by (prove_unique_tac lemma 1);
   120 qed "unique_NA";
   120 qed "unique_NA";
   121 
   121 
   122 
   122 
   128 
   128 
   129 
   129 
   130 (*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure*)
   130 (*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure*)
   131 goal thy 
   131 goal thy 
   132  "!!evs. [| Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) : set evs;         \
   132  "!!evs. [| Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) : set evs;         \
   133 \           A ~: lost;  B ~: lost;  evs : ns_public |]                        \
   133 \           A ~: bad;  B ~: bad;  evs : ns_public |]                        \
   134 \        ==>  Nonce NA ~: analz (sees Spy evs)";
   134 \        ==>  Nonce NA ~: analz (spies evs)";
   135 by (etac rev_mp 1);
   135 by (etac rev_mp 1);
   136 by (analz_induct_tac 1);
   136 by (analz_induct_tac 1);
   137 (*NS3*)
   137 (*NS3*)
   138 by (blast_tac (!claset addDs  [Says_imp_sees_Spy RS parts.Inj]
   138 by (blast_tac (!claset addDs  [Says_imp_spies RS parts.Inj]
   139                        addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 4);
   139                        addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 4);
   140 (*NS2*)
   140 (*NS2*)
   141 by (blast_tac (!claset addSEs [MPair_parts]
   141 by (blast_tac (!claset addSEs [MPair_parts]
   142 		       addDs  [Says_imp_sees_Spy RS parts.Inj,
   142 		       addDs  [Says_imp_spies RS parts.Inj,
   143 			       parts.Body, unique_NA]) 3);
   143 			       parts.Body, unique_NA]) 3);
   144 (*NS1*)
   144 (*NS1*)
   145 by (blast_tac (!claset addSEs sees_Spy_partsEs
   145 by (blast_tac (!claset addSEs spies_partsEs
   146                        addIs  [impOfSubs analz_subset_parts]) 2);
   146                        addIs  [impOfSubs analz_subset_parts]) 2);
   147 (*Fake*)
   147 (*Fake*)
   148 by (spy_analz_tac 1);
   148 by (spy_analz_tac 1);
   149 qed "Spy_not_see_NA";
   149 qed "Spy_not_see_NA";
   150 
   150 
   153   to start a run, then B has sent message 2.*)
   153   to start a run, then B has sent message 2.*)
   154 goal thy 
   154 goal thy 
   155  "!!evs. [| Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs;  \
   155  "!!evs. [| Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs;  \
   156 \           Says B' A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|})   \
   156 \           Says B' A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|})   \
   157 \             : set evs;                                                \
   157 \             : set evs;                                                \
   158 \           A ~: lost;  B ~: lost;  evs : ns_public |]                  \
   158 \           A ~: bad;  B ~: bad;  evs : ns_public |]                  \
   159 \        ==> Says B A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|})   \
   159 \        ==> Says B A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|})   \
   160 \              : set evs";
   160 \              : set evs";
   161 by (etac rev_mp 1);
   161 by (etac rev_mp 1);
   162 (*prepare induction over Crypt (pubK A) {|NA,NB,B|} : parts H*)
   162 (*prepare induction over Crypt (pubK A) {|NA,NB,B|} : parts H*)
   163 by (etac (Says_imp_sees_Spy RS parts.Inj RS rev_mp) 1);
   163 by (etac (Says_imp_spies RS parts.Inj RS rev_mp) 1);
   164 by (etac ns_public.induct 1);
   164 by (etac ns_public.induct 1);
   165 by (ALLGOALS Asm_simp_tac);
   165 by (ALLGOALS Asm_simp_tac);
   166 (*NS1*)
   166 (*NS1*)
   167 by (blast_tac (!claset addSEs sees_Spy_partsEs) 2);
   167 by (blast_tac (!claset addSEs spies_partsEs) 2);
   168 (*Fake*)
   168 (*Fake*)
   169 by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert]
   169 by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert]
   170                        addDs  [Spy_not_see_NA, 
   170                        addDs  [Spy_not_see_NA, 
   171 			       impOfSubs analz_subset_parts]) 1);
   171 			       impOfSubs analz_subset_parts]) 1);
   172 qed "A_trusts_NS2";
   172 qed "A_trusts_NS2";
   173 
   173 
   174 (*If the encrypted message appears then it originated with Alice in NS1*)
   174 (*If the encrypted message appears then it originated with Alice in NS1*)
   175 goal thy 
   175 goal thy 
   176  "!!evs. [| Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees Spy evs); \
   176  "!!evs. [| Crypt (pubK B) {|Nonce NA, Agent A|} : parts (spies evs); \
   177 \           Nonce NA ~: analz (sees Spy evs);                 \
   177 \           Nonce NA ~: analz (spies evs);                 \
   178 \           evs : ns_public |]                                     \
   178 \           evs : ns_public |]                             \
   179 \   ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs";
   179 \   ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs";
   180 by (etac rev_mp 1);
   180 by (etac rev_mp 1);
   181 by (etac rev_mp 1);
   181 by (etac rev_mp 1);
   182 by (parts_induct_tac 1);
   182 by (parts_induct_tac 1);
   183 by (Fake_parts_insert_tac 1);
   183 by (Fake_parts_insert_tac 1);
   189 
   189 
   190 (*Unicity for NS2: nonce NB identifies nonce NA and agents A, B 
   190 (*Unicity for NS2: nonce NB identifies nonce NA and agents A, B 
   191   [unicity of B makes Lowe's fix work]
   191   [unicity of B makes Lowe's fix work]
   192   [proof closely follows that for unique_NA] *)
   192   [proof closely follows that for unique_NA] *)
   193 goal thy 
   193 goal thy 
   194  "!!evs. [| Nonce NB ~: analz (sees Spy evs);  evs : ns_public |]      \
   194  "!!evs. [| Nonce NB ~: analz (spies evs);  evs : ns_public |]      \
   195 \ ==> EX A' NA' B'. ALL A NA B.                                             \
   195 \ ==> EX A' NA' B'. ALL A NA B.                                             \
   196 \      Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}                       \
   196 \      Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|} : parts (spies evs) \
   197 \        : parts (sees Spy evs)  -->  A=A' & NA=NA' & B=B'";
   197 \         -->  A=A' & NA=NA' & B=B'";
   198 by (etac rev_mp 1);
   198 by (etac rev_mp 1);
   199 by (parts_induct_tac 1);
   199 by (parts_induct_tac 1);
   200 by (ALLGOALS
   200 by (ALLGOALS
   201     (asm_simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_sees])));
   201     (asm_simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_spies])));
   202 (*NS2*)
   202 (*NS2*)
   203 by (expand_case_tac "NB = ?y" 2 THEN blast_tac (!claset addSEs partsEs) 2);
   203 by (expand_case_tac "NB = ?y" 2 THEN blast_tac (!claset addSEs partsEs) 2);
   204 (*Fake*)
   204 (*Fake*)
   205 by (step_tac (!claset addSIs [analz_insertI]) 1);
   205 by (step_tac (!claset addSIs [analz_insertI]) 1);
   206 by (ex_strip_tac 1);
   206 by (ex_strip_tac 1);
   207 by (Fake_parts_insert_tac 1);
   207 by (Fake_parts_insert_tac 1);
   208 val lemma = result();
   208 val lemma = result();
   209 
   209 
   210 goal thy 
   210 goal thy 
   211  "!!evs. [| Crypt(pubK A)  {|Nonce NA, Nonce NB, Agent B|}   \
   211  "!!evs. [| Crypt(pubK A)  {|Nonce NA, Nonce NB, Agent B|}   \
   212 \             : parts(sees Spy evs);                         \
   212 \             : parts(spies evs);                         \
   213 \           Crypt(pubK A') {|Nonce NA', Nonce NB, Agent B'|} \
   213 \           Crypt(pubK A') {|Nonce NA', Nonce NB, Agent B'|} \
   214 \             : parts(sees Spy evs);                         \
   214 \             : parts(spies evs);                         \
   215 \           Nonce NB ~: analz (sees Spy evs);                \
   215 \           Nonce NB ~: analz (spies evs);                \
   216 \           evs : ns_public |]                               \
   216 \           evs : ns_public |]                               \
   217 \        ==> A=A' & NA=NA' & B=B'";
   217 \        ==> A=A' & NA=NA' & B=B'";
   218 by (prove_unique_tac lemma 1);
   218 by (prove_unique_tac lemma 1);
   219 qed "unique_NB";
   219 qed "unique_NB";
   220 
   220 
   221 
   221 
   222 (*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)
   222 (*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)
   223 goal thy 
   223 goal thy 
   224  "!!evs. [| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
   224  "!!evs. [| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
   225 \             : set evs;                                              \
   225 \             : set evs;                                              \
   226 \           A ~: lost;  B ~: lost;  evs : ns_public |]                \
   226 \           A ~: bad;  B ~: bad;  evs : ns_public |]                \
   227 \ ==> Nonce NB ~: analz (sees Spy evs)";
   227 \ ==> Nonce NB ~: analz (spies evs)";
   228 by (etac rev_mp 1);
   228 by (etac rev_mp 1);
   229 by (analz_induct_tac 1);
   229 by (analz_induct_tac 1);
   230 (*NS3*)
   230 (*NS3*)
   231 by (blast_tac (!claset addDs [Says_imp_sees_Spy RS parts.Inj, unique_NB]) 4);
   231 by (blast_tac (!claset addDs [Says_imp_spies RS parts.Inj, unique_NB]) 4);
   232 (*NS1*)
   232 (*NS1*)
   233 by (blast_tac (!claset addSEs sees_Spy_partsEs) 2);
   233 by (blast_tac (!claset addSEs spies_partsEs) 2);
   234 (*Fake*)
   234 (*Fake*)
   235 by (spy_analz_tac 1);
   235 by (spy_analz_tac 1);
   236 (*NS2*)
   236 (*NS2*)
   237 by (Step_tac 1);
   237 by (Step_tac 1);
   238 by (blast_tac (!claset addSEs sees_Spy_partsEs) 3);
   238 by (blast_tac (!claset addSEs spies_partsEs) 3);
   239 by (blast_tac (!claset addSDs [Says_imp_sees_Spy RS parts.Inj]
   239 by (blast_tac (!claset addSDs [Says_imp_spies RS parts.Inj]
   240                        addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 2);
   240                        addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 2);
   241 by (blast_tac (!claset addSIs [impOfSubs analz_subset_parts]) 1);
   241 by (blast_tac (!claset addSIs [impOfSubs analz_subset_parts]) 1);
   242 qed "Spy_not_see_NB";
   242 qed "Spy_not_see_NB";
   243 
   243 
   244 
   244 
   246   in message 2, then A has sent message 3.*)
   246   in message 2, then A has sent message 3.*)
   247 goal thy 
   247 goal thy 
   248  "!!evs. [| Says B A  (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
   248  "!!evs. [| Says B A  (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
   249 \             : set evs;                                               \
   249 \             : set evs;                                               \
   250 \           Says A' B (Crypt (pubK B) (Nonce NB)): set evs;            \
   250 \           Says A' B (Crypt (pubK B) (Nonce NB)): set evs;            \
   251 \           A ~: lost;  B ~: lost;  evs : ns_public |]                 \
   251 \           A ~: bad;  B ~: bad;  evs : ns_public |]                 \
   252 \        ==> Says A B (Crypt (pubK B) (Nonce NB)) : set evs";
   252 \        ==> Says A B (Crypt (pubK B) (Nonce NB)) : set evs";
   253 by (etac rev_mp 1);
   253 by (etac rev_mp 1);
   254 (*prepare induction over Crypt (pubK B) NB : parts H*)
   254 (*prepare induction over Crypt (pubK B) NB : parts H*)
   255 by (etac (Says_imp_sees_Spy RS parts.Inj RS rev_mp) 1);
   255 by (etac (Says_imp_spies RS parts.Inj RS rev_mp) 1);
   256 by (parts_induct_tac 1);
   256 by (parts_induct_tac 1);
   257 (*NS1*)
   257 (*NS1*)
   258 by (blast_tac (!claset addSEs sees_Spy_partsEs) 2);
   258 by (blast_tac (!claset addSEs spies_partsEs) 2);
   259 (*Fake*)
   259 (*Fake*)
   260 by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert]
   260 by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert]
   261                        addDs  [Spy_not_see_NB, 
   261                        addDs  [Spy_not_see_NB, 
   262 			       impOfSubs analz_subset_parts]) 1);
   262 			       impOfSubs analz_subset_parts]) 1);
   263 (*NS3; not clear why blast_tac needs to be preceeded by Step_tac*)
   263 (*NS3; not clear why blast_tac needs to be preceeded by Step_tac*)
   264 by (Step_tac 1);
   264 by (Step_tac 1);
   265 by (blast_tac (!claset addDs [Says_imp_sees_Spy RS parts.Inj,
   265 by (blast_tac (!claset addDs [Says_imp_spies RS parts.Inj,
   266 			      Spy_not_see_NB, unique_NB]) 1);
   266 			      Spy_not_see_NB, unique_NB]) 1);
   267 qed "B_trusts_NS3";
   267 qed "B_trusts_NS3";
   268 
   268 
   269 
   269 
   270 (**** Overall guarantee for B*)
   270 (**** Overall guarantee for B*)
   271 
   271 
   272 (*Matches only NS2, not NS1 (or NS3)*)
   272 (*Matches only NS2, not NS1 (or NS3)*)
   273 val Says_imp_sees_Spy' = 
   273 val Says_imp_spies' = 
   274     read_instantiate [("X","Crypt ?K {|?XX,?YY,?ZZ|}")] Says_imp_sees_Spy;
   274     read_instantiate [("X","Crypt ?K {|?XX,?YY,?ZZ|}")] Says_imp_spies;
   275 
   275 
   276 
   276 
   277 (*If B receives NS3 and the nonce NB agrees with the nonce he joined with
   277 (*If B receives NS3 and the nonce NB agrees with the nonce he joined with
   278   NA, then A initiated the run using NA.  SAME proof as B_trusts_NS3!*)
   278   NA, then A initiated the run using NA.  SAME proof as B_trusts_NS3!*)
   279 goal thy 
   279 goal thy 
   280  "!!evs. [| Says B A  (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
   280  "!!evs. [| Says B A  (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
   281 \             : set evs;                                               \
   281 \             : set evs;                                               \
   282 \           Says A' B (Crypt (pubK B) (Nonce NB)): set evs;            \
   282 \           Says A' B (Crypt (pubK B) (Nonce NB)): set evs;            \
   283 \           A ~: lost;  B ~: lost;  evs : ns_public |]                 \
   283 \           A ~: bad;  B ~: bad;  evs : ns_public |]                 \
   284 \    ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs";
   284 \    ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs";
   285 by (etac rev_mp 1);
   285 by (etac rev_mp 1);
   286 (*prepare induction over Crypt (pubK B) {|NB|} : parts H*)
   286 (*prepare induction over Crypt (pubK B) {|NB|} : parts H*)
   287 by (etac (Says_imp_sees_Spy RS parts.Inj RS rev_mp) 1);
   287 by (etac (Says_imp_spies RS parts.Inj RS rev_mp) 1);
   288 by (etac ns_public.induct 1);
   288 by (etac ns_public.induct 1);
   289 by (ALLGOALS Asm_simp_tac);
   289 by (ALLGOALS Asm_simp_tac);
   290 (*Fake, NS2, NS3*)
   290 (*Fake, NS2, NS3*)
   291 (*NS1*)
   291 (*NS1*)
   292 by (blast_tac (!claset addSEs sees_Spy_partsEs) 2);
   292 by (blast_tac (!claset addSEs spies_partsEs) 2);
   293 (*Fake*)
   293 (*Fake*)
   294 by (REPEAT_FIRST (resolve_tac [impI, conjI]));
   294 by (REPEAT_FIRST (resolve_tac [impI, conjI]));
   295 by (Blast_tac 1);
   295 by (Blast_tac 1);
   296 by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
   296 by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
   297 by (blast_tac (!claset addSIs [disjI2]
   297 by (blast_tac (!claset addSIs [disjI2]
   298                       addDs [impOfSubs analz_subset_parts,
   298                       addDs [impOfSubs analz_subset_parts,
   299                              impOfSubs Fake_parts_insert]) 1);
   299                              impOfSubs Fake_parts_insert]) 1);
   300 (*NS3*)
   300 (*NS3*)
   301 by (Step_tac 1);
   301 by (Step_tac 1);
   302 by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
   302 by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
   303 by (blast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
   303 by (blast_tac (!claset addSDs [Says_imp_spies' RS parts.Inj]
   304                        addDs  [unique_NB]) 1);
   304                        addDs  [unique_NB]) 1);
   305 qed "B_trusts_protocol";
   305 qed "B_trusts_protocol";
   306 
   306