src/HOL/Auth/NS_Public.ML
author paulson
Tue, 07 Jan 1997 10:18:20 +0100
changeset 2480 f9be937df511
parent 2451 ce85a2aafc7a
child 2497 47de509bdd55
permissions -rw-r--r--
Tidied up the unicity proofs

(*  Title:      HOL/Auth/NS_Public
    ID:         $Id$
    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 identify in round 2).
*)

open NS_Public;

proof_timing:=true;
HOL_quantifiers := false;
Pretty.setdepth 20;

AddIffs [Spy_in_lost];

(*Replacing the variable by a constant improves search speed by 50%!*)
val Says_imp_sees_Spy' = 
    read_instantiate_sg (sign_of thy) [("lost","lost")] Says_imp_sees_Spy;


(*A "possibility property": there are traces that reach the end*)
goal thy 
 "!!A B. A ~= B ==> EX NB. EX evs: ns_public.               \
\                     Says A B (Crypt (pubK B) (Nonce NB)) : set_of_list evs";
by (REPEAT (resolve_tac [exI,bexI] 1));
by (rtac (ns_public.Nil RS ns_public.NS1 RS ns_public.NS2 RS ns_public.NS3) 2);
by (ALLGOALS (simp_tac (!simpset setsolver safe_solver)));
by (REPEAT_FIRST (resolve_tac [refl, conjI]));
by (REPEAT_FIRST (fast_tac (!claset addss (!simpset setsolver safe_solver))));
result();


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

(*Nobody sends themselves messages*)
goal thy "!!evs. evs : ns_public ==> ALL A X. Says A A X ~: set_of_list evs";
by (etac ns_public.induct 1);
by (Auto_tac());
qed_spec_mp "not_Says_to_self";
Addsimps [not_Says_to_self];
AddSEs   [not_Says_to_self RSN (2, rev_notE)];


(*For proving the easier theorems about X ~: parts (sees lost Spy evs) *)
fun parts_induct_tac i = SELECT_GOAL
    (DETERM (etac ns_public.induct 1 THEN 
             (*Fake message*)
             TRY (best_tac (!claset addDs [impOfSubs analz_subset_parts,
                                           impOfSubs Fake_parts_insert]
                                    addss (!simpset)) 2)) THEN
     (*Base case*)
     fast_tac (!claset addss (!simpset)) 1 THEN
     ALLGOALS Asm_simp_tac) i;

(** Theorems of the form X ~: parts (sees lost Spy evs) imply that NOBODY
    sends messages containing X! **)

(*Spy never sees another agent's private key! (unless it's lost at start)*)
goal thy 
 "!!evs. evs : ns_public \
\        ==> (Key (priK A) : parts (sees lost Spy evs)) = (A : lost)";
by (parts_induct_tac 1);
by (Auto_tac());
qed "Spy_see_priK";
Addsimps [Spy_see_priK];

goal thy 
 "!!evs. evs : ns_public \
\        ==> (Key (priK A) : analz (sees lost Spy evs)) = (A : lost)";
by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset));
qed "Spy_analz_priK";
Addsimps [Spy_analz_priK];

goal thy  "!!A. [| Key (priK A) : parts (sees lost Spy evs);       \
\                  evs : ns_public |] ==> A:lost";
by (fast_tac (!claset addDs [Spy_see_priK]) 1);
qed "Spy_see_priK_D";

bind_thm ("Spy_analz_priK_D", analz_subset_parts RS subsetD RS Spy_see_priK_D);
AddSDs [Spy_see_priK_D, Spy_analz_priK_D];


(*** Future nonces can't be seen or used! ***)

goal thy "!!i. evs : ns_public ==>        \
\             length evs <= i --> Nonce (newN i) ~: parts (sees lost Spy evs)";
by (parts_induct_tac 1);
by (REPEAT_FIRST (fast_tac (!claset 
                              addSEs partsEs
                              addSDs  [Says_imp_sees_Spy' RS parts.Inj]
                              addEs [leD RS notE]
			      addDs  [impOfSubs analz_subset_parts,
                                      impOfSubs parts_insert_subset_Un,
                                      Suc_leD]
                              addss (!simpset))));
qed_spec_mp "new_nonces_not_seen";
Addsimps [new_nonces_not_seen];

val nonce_not_seen_now = le_refl RSN (2, new_nonces_not_seen) RSN (2,rev_notE);

fun analz_induct_tac i = 
    etac ns_public.induct i	THEN
    ALLGOALS (asm_simp_tac 
	      (!simpset addsimps [not_parts_not_analz]
                        setloop split_tac [expand_if]));

(**** 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.)*)
goal thy 
 "!!evs. evs : ns_public                       \
\ ==> Nonce NA ~: analz (sees lost Spy evs) -->           \
\     Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs) --> \
\     Crypt (pubK C) {|NA', Nonce NA, Agent D|} ~: parts (sees lost Spy evs)";
by (analz_induct_tac 1);
(*NS3*)
by (fast_tac (!claset addSEs partsEs
                      addEs  [nonce_not_seen_now]) 4);
(*NS2*)
by (fast_tac (!claset addSEs partsEs
                      addEs  [nonce_not_seen_now]) 3);
(*Fake*)
by (best_tac (!claset addIs [analz_insertI]
		      addDs [impOfSubs analz_subset_parts,
			     impOfSubs Fake_parts_insert]
	              addss (!simpset)) 2);
(*Base*)
by (fast_tac (!claset addss (!simpset)) 1);
bind_thm ("no_nonce_NS1_NS2", result() RSN (2, rev_mp) RSN (2, rev_mp));


(*Unicity for NS1: nonce NA identifies agents A and B*)
goal thy 
 "!!evs. evs : ns_public                                                    \
\ ==> Nonce NA ~: analz (sees lost Spy evs) -->                             \
\     (EX A' B'. ALL A B.                                                   \
\      Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs) --> \
\      A=A' & B=B')";
by (analz_induct_tac 1);
(*NS1*)
by (expand_case_tac "NA = ?y" 3 THEN
    REPEAT (fast_tac (!claset addSEs (nonce_not_seen_now::partsEs)) 3));
(*Base*)
by (fast_tac (!claset addss (!simpset)) 1);
(*Fake*)
by (simp_tac (!simpset addsimps [all_conj_distrib,parts_insert_sees]) 1);
by (step_tac (!claset addSIs [analz_insertI]) 1);
by (ex_strip_tac 1);
by (best_tac (!claset delrules [conjI]
		      addSDs [impOfSubs Fake_parts_insert]
	              addDs  [impOfSubs analz_subset_parts]
                      addss (!simpset)) 1);
val lemma = result();

goal thy 
 "!!evs. [| Crypt(pubK B)  {|Nonce NA, Agent A|}  : parts(sees lost Spy evs); \
\           Crypt(pubK B') {|Nonce NA, Agent A'|} : parts(sees lost Spy evs); \
\           Nonce NA ~: analz (sees lost Spy evs);                            \
\           evs : ns_public |]                                                \
\        ==> A=A' & B=B'";
by (prove_unique_tac lemma 1);
qed "unique_NA";


(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure*)
goal thy 
 "!!evs. [| A ~: lost;  B ~: lost;  evs : ns_public |]   \
\ ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs \
\     --> Nonce NA ~: analz (sees lost Spy evs)";
by (analz_induct_tac 1);
(*NS3*)
by (fast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
                      addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 4);
(*NS1*)
by (fast_tac (!claset addSEs partsEs
                      addSDs [new_nonces_not_seen, 
			      Says_imp_sees_Spy' RS parts.Inj]) 2);
(*Fake*)
by (REPEAT_FIRST spy_analz_tac);
(*NS2*)
by (Step_tac 1);
by (fast_tac (!claset addSEs partsEs
                      addSDs [new_nonces_not_seen, 
			      Says_imp_sees_Spy' RS parts.Inj]) 2);
by (fast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
                      addDs  [unique_NA]) 1);
bind_thm ("Spy_not_see_NA", result() RSN (2, rev_mp));


(*Authentication for A: if she receives message 2 and has used NA
  to start a run, then B has sent message 2.*)
goal thy 
 "!!evs. [| A ~: lost;  B ~: lost;  evs : ns_public |]       \
\ ==> Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}         \
\                                : parts (sees lost Spy evs) \
\     --> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs \
\     --> Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})         \
\           : set_of_list evs";
by (etac ns_public.induct 1);
by (ALLGOALS Asm_simp_tac);
(*NS1*)
by (fast_tac (!claset addSEs partsEs
                      addSDs [new_nonces_not_seen, 
			      Says_imp_sees_Spy' RS parts.Inj]) 2);
(*Fake*)
by (REPEAT_FIRST (resolve_tac [impI, conjI]));
by (fast_tac (!claset addss (!simpset)) 1);
by (forward_tac [Spy_not_see_NA] 1 THEN REPEAT (assume_tac 1));
by (best_tac (!claset addSIs [disjI2]
		      addSDs [impOfSubs Fake_parts_insert]
		      addDs  [impOfSubs analz_subset_parts]
	              addss (!simpset)) 1);
qed_spec_mp "NA_decrypt_imp_B_msg";

(*Corollary: if A receives B's message NS2 and the nonce NA agrees
  then that message really originated with B.*)
goal thy 
 "!!evs. [| Says B' A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
\             : set_of_list evs;\
\           Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs;\
\           A ~: lost;  B ~: lost;  evs : ns_public |]  \
\        ==> Says B A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
\              : set_of_list evs";
by (fast_tac (!claset addSIs [NA_decrypt_imp_B_msg]
                      addEs  partsEs
                      addDs  [Says_imp_sees_Spy' RS parts.Inj]) 1);
qed "A_trusts_NS2";

(*If the encrypted message appears then it originated with Alice in NS1*)
goal thy 
 "!!evs. evs : ns_public                   \
\    ==> Nonce NA ~: analz (sees lost Spy evs) --> \
\        Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs) --> \
\        Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs";
by (analz_induct_tac 1);
(*Fake*)
by (best_tac (!claset addSIs [disjI2]
		      addSDs [impOfSubs Fake_parts_insert]
		      addIs  [analz_insertI]
		      addDs  [impOfSubs analz_subset_parts]
	              addss (!simpset)) 2);
(*Base*)
by (fast_tac (!claset addss (!simpset)) 1);
qed_spec_mp "B_trusts_NS1";



(**** 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] *)
goal thy 
 "!!evs. evs : ns_public                                                    \
\ ==> Nonce NB ~: analz (sees lost Spy evs) -->                             \
\     (EX A' NA' B'. ALL A NA B.                                            \
\      Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}   \
\        : parts (sees lost Spy evs)  -->  A=A' & NA=NA' & B=B')";
by (analz_induct_tac 1);
(*NS2*)
by (expand_case_tac "NB = ?y" 3 THEN
    REPEAT (fast_tac (!claset addSEs (nonce_not_seen_now::partsEs)) 3));
(*Base*)
by (fast_tac (!claset addss (!simpset)) 1);
(*Fake*)
by (simp_tac (!simpset addsimps [all_conj_distrib,parts_insert_sees]) 1);
by (step_tac (!claset addSIs [analz_insertI]) 1);
by (ex_strip_tac 1);
by (best_tac (!claset delrules [conjI]
	              addSDs [impOfSubs Fake_parts_insert]
	              addDs  [impOfSubs analz_subset_parts] 
	              addss (!simpset)) 1);
val lemma = result();

goal thy 
 "!!evs. [| Crypt(pubK A)  {|Nonce NA, Nonce NB, Agent B|}   \
\             : parts(sees lost Spy evs);                    \
\           Crypt(pubK A') {|Nonce NA', Nonce NB, Agent B'|} \
\             : parts(sees lost Spy evs);                    \
\           Nonce NB ~: analz (sees lost Spy evs);           \
\           evs : ns_public |]                               \
\        ==> A=A' & NA=NA' & B=B'";
by (prove_unique_tac lemma 1);
qed "unique_NB";


(*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)
goal thy 
 "!!evs. [| A ~: lost;  B ~: lost;  evs : ns_public |]   \
\ ==> Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
\       : set_of_list evs \
\     --> Nonce NB ~: analz (sees lost Spy evs)";
by (analz_induct_tac 1);
(*NS3*)
by (fast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
                      addDs  [unique_NB]) 4);
(*NS1*)
by (fast_tac (!claset addSEs partsEs
                      addSDs [new_nonces_not_seen, 
			      Says_imp_sees_Spy' RS parts.Inj]) 2);
(*Fake*)
by (REPEAT_FIRST spy_analz_tac);
(*NS2*)
by (Step_tac 1);
by (fast_tac (!claset addSEs partsEs
                      addSDs [new_nonces_not_seen, 
                              Says_imp_sees_Spy' RS parts.Inj]) 2);
by (fast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
                      addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 1);
bind_thm ("Spy_not_see_NB", result() RSN (2, rev_mp));


(*Matches only NS2, not NS1 (or NS3)*)
val Says_imp_sees_Spy'' = 
    read_instantiate [("X","Crypt ?K {|?XX,?YY,?ZZ|}")] Says_imp_sees_Spy';


(*Authentication for B: if he receives message 3 and has used NB
  in message 2, then A has sent message 3.*)
goal thy 
 "!!evs. [| A ~: lost;  B ~: lost;  evs : ns_public |]       \
\ ==> Crypt (pubK B) (Nonce NB) : parts (sees lost Spy evs) \
\     --> Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
\       : set_of_list evs \
\     --> Says A B (Crypt (pubK B) (Nonce NB)) : set_of_list evs";
by (etac ns_public.induct 1);
by (ALLGOALS Asm_simp_tac);
(*NS1*)
by (fast_tac (!claset addSEs partsEs
                      addSDs [new_nonces_not_seen, 
			      Says_imp_sees_Spy' RS parts.Inj]) 2);
(*Fake*)
by (REPEAT_FIRST (resolve_tac [impI, conjI]));
by (fast_tac (!claset addss (!simpset)) 1);
by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
by (best_tac (!claset addSIs [disjI2]
		      addSDs [impOfSubs Fake_parts_insert]
	              addDs  [impOfSubs analz_subset_parts] 
	              addss (!simpset)) 1);
(*NS3*)
by (Step_tac 1);
by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
by (best_tac (!claset addSDs [Says_imp_sees_Spy'' RS parts.Inj]
	              addDs  [unique_NB]) 1);
qed_spec_mp "NB_decrypt_imp_A_msg";

(*Corollary: if B receives message NS3 and the nonce NB agrees,
  then that message really originated with A.*)
goal thy 
 "!!evs. [| Says A' B (Crypt (pubK B) (Nonce NB)): set_of_list evs;    \
\           Says B A  (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
\             : set_of_list evs;                                       \
\           A ~: lost;  B ~: lost;  evs : ns_public |]                 \
\        ==> Says A B (Crypt (pubK B) (Nonce NB)) : set_of_list evs";
by (fast_tac (!claset addSIs [NB_decrypt_imp_A_msg]
                      addEs  partsEs
                      addDs  [Says_imp_sees_Spy' RS parts.Inj]) 1);
qed "B_trusts_NS3";


(**** Overall guarantee for B*)

(*If B receives NS3 and the nonce NB agrees with the nonce he joined with
  NA, then A initiated the run using NA.
  SAME PROOF AS NB_decrypt_imp_A_msg*)
goal thy 
 "!!evs. [| A ~: lost;  B ~: lost;  evs : ns_public |]       \
\ ==> Crypt (pubK B) (Nonce NB) : parts (sees lost Spy evs) \
\     --> Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
\       : set_of_list evs \
\     --> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs";
by (etac ns_public.induct 1);
by (ALLGOALS Asm_simp_tac);
(*Fake, NS2, NS3*)
(*NS1*)
by (fast_tac (!claset addSEs partsEs
                      addSDs [new_nonces_not_seen, 
			      Says_imp_sees_Spy' RS parts.Inj]) 2);
(*Fake*)
by (REPEAT_FIRST (resolve_tac [impI, conjI]));
by (fast_tac (!claset addss (!simpset)) 1);
by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
by (best_tac (!claset addSIs [disjI2]
		      addDs [impOfSubs analz_subset_parts,
			     impOfSubs Fake_parts_insert]
	              addss (!simpset)) 1);
(*NS3*)
by (Step_tac 1);
by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1));
by (best_tac (!claset addSDs [Says_imp_sees_Spy'' RS parts.Inj]
	              addDs  [unique_NB]) 1);
val lemma = result() RSN (2, rev_mp) RSN (2, rev_mp);

goal thy 
 "!!evs. [| Says A' B (Crypt (pubK B) (Nonce NB)): set_of_list evs;    \
\           Says B A  (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
\             : set_of_list evs;                                       \
\           A ~: lost;  B ~: lost;  evs : ns_public |]                 \
\    ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs";
by (fast_tac (!claset addSIs [lemma]
                      addEs  partsEs
                      addDs  [Says_imp_sees_Spy' RS parts.Inj]) 1);
qed_spec_mp "B_trusts_protocol";