conversion of Auth/TLS to Isar script
authorpaulson
Tue, 08 May 2001 15:56:57 +0200
changeset 11287 0103ee3082bf
parent 11286 5116d92c6a83
child 11288 7fe6593133d4
conversion of Auth/TLS to Isar script
src/HOL/Auth/Public.thy
src/HOL/Auth/TLS.ML
src/HOL/Auth/TLS.thy
src/HOL/IsaMakefile
--- a/src/HOL/Auth/Public.thy	Mon May 07 19:19:41 2001 +0200
+++ b/src/HOL/Auth/Public.thy	Tue May 08 15:56:57 2001 +0200
@@ -37,6 +37,9 @@
 
 use "Public_lemmas.ML"
 
+lemma [simp]: "K \\<in> symKeys ==> invKey K = K"
+by (simp add: symKeys_def)
+
 (*Specialized methods*)
 
 method_setup possibility = {*
--- a/src/HOL/Auth/TLS.ML	Mon May 07 19:19:41 2001 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,660 +0,0 @@
-(*  Title:      HOL/Auth/TLS
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
-
-Protocol goals: 
-* M, serverK(NA,NB,M) and clientK(NA,NB,M) will be known only to the two
-     parties (though A is not necessarily authenticated).
-
-* B upon receiving CertVerify knows that A is present (But this
-    message is optional!)
-
-* A upon receiving ServerFinished knows that B is present
-
-* Each party who has received a FINISHED message can trust that the other
-  party agrees on all message components, including PA and PB (thus foiling
-  rollback attacks).
-*)
-
-AddDs  [Says_imp_knows_Spy RS parts.Inj, parts.Body];
-AddDs  [impOfSubs analz_subset_parts, impOfSubs Fake_parts_insert];
-
-(*Automatically unfold the definition of "certificate"*)
-Addsimps [certificate_def];
-
-(*Injectiveness of key-generating functions*)
-AddIffs [inj_PRF RS inj_eq, inj_sessionK RS inj_eq];
-
-(* invKey(sessionK x) = sessionK x*)
-Addsimps [isSym_sessionK, rewrite_rule [symKeys_def] isSym_sessionK];
-
-
-(*** clientK and serverK make symmetric keys; no clashes with pubK or priK ***)
-
-Goal "pubK A \\<noteq> sessionK arg";
-by (simp_tac (simpset() addsimps [symKeys_neq_imp_neq]) 1);
-qed "pubK_neq_sessionK";
-
-Goal "priK A \\<noteq> sessionK arg";
-by (simp_tac (simpset() addsimps [symKeys_neq_imp_neq]) 1);
-qed "priK_neq_sessionK";
-
-val keys_distinct = [pubK_neq_sessionK, priK_neq_sessionK];
-AddIffs (keys_distinct @ (keys_distinct RL [not_sym]));
-
-
-(**** Protocol Proofs ****)
-
-(*Possibility properties state that some traces run the protocol to the end.
-  Four paths and 12 rules are considered.*)
-
-
-(** These proofs assume that the Nonce_supply nonces 
-	(which have the form  @ N. Nonce N \\<notin> used evs)
-    lie outside the range of PRF.  It seems reasonable, but as it is needed
-    only for the possibility theorems, it is not taken as an axiom.
-**)
-
-
-(*Possibility property ending with ClientAccepts.*)
-Goal "[| \\<forall>evs. (@ N. Nonce N \\<notin> used evs) \\<notin> range PRF;  \
-\        A \\<noteq> B |]            \
-\     ==> \\<exists>SID M. \\<exists>evs \\<in> tls.    \
-\          Notes A {|Number SID, Agent A, Agent B, Nonce M|} \\<in> set evs";
-by (REPEAT (resolve_tac [exI,bexI] 1));
-by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.Certificate RS
-	  tls.ClientKeyExch RS tls.ClientFinished RS tls.ServerFinished RS
-	  tls.ClientAccepts) 2);
-by possibility_tac;
-by (REPEAT (Blast_tac 1));
-result();
-
-(*And one for ServerAccepts.  Either FINISHED message may come first.*)
-Goal "[| \\<forall>evs. (@ N. Nonce N \\<notin> used evs) \\<notin> range PRF;  \
-\        A \\<noteq> B |]                        \
-\     ==> \\<exists>SID NA PA NB PB M. \\<exists>evs \\<in> tls.    \
-\          Notes B {|Number SID, Agent A, Agent B, Nonce M|} \\<in> set evs";
-by (REPEAT (resolve_tac [exI,bexI] 1));
-by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.Certificate RS
-	  tls.ClientKeyExch RS tls.ServerFinished RS tls.ClientFinished RS
-	  tls.ServerAccepts) 2);
-by possibility_tac;
-by (REPEAT (Blast_tac 1));
-result();
-
-(*Another one, for CertVerify (which is optional)*)
-Goal "[| \\<forall>evs. (@ N. Nonce N \\<notin> used evs) \\<notin> range PRF;  \
-\        A \\<noteq> B |]                       \
-\  ==> \\<exists>NB PMS. \\<exists>evs \\<in> tls.   \
-\  Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|})) \\<in> set evs";
-by (REPEAT (resolve_tac [exI,bexI] 1));
-by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.Certificate RS
-	  tls.ClientKeyExch RS tls.CertVerify) 2);
-by possibility_tac;
-by (REPEAT (Blast_tac 1));
-result();
-
-(*Another one, for session resumption (both ServerResume and ClientResume) *)
-Goal "[| evs0 \\<in> tls;     \
-\        Notes A {|Number SID, Agent A, Agent B, Nonce M|} \\<in> set evs0; \
-\        Notes B {|Number SID, Agent A, Agent B, Nonce M|} \\<in> set evs0; \
-\        \\<forall>evs. (@ N. Nonce N \\<notin> used evs) \\<notin> range PRF;  \
-\        A \\<noteq> B |] \
-\     ==> \\<exists>NA PA NB PB X. \\<exists>evs \\<in> tls.    \
-\           X = Hash{|Number SID, Nonce M,             \
-\                     Nonce NA, Number PA, Agent A,      \
-\                     Nonce NB, Number PB, Agent B|}  &  \
-\           Says A B (Crypt (clientK(NA,NB,M)) X) \\<in> set evs  &  \
-\           Says B A (Crypt (serverK(NA,NB,M)) X) \\<in> set evs";
-by (REPEAT (resolve_tac [exI,bexI] 1));
-by (etac (tls.ClientHello RS tls.ServerHello RS tls.ServerResume RS 
-	  tls.ClientResume) 2);
-by possibility_tac;
-by (REPEAT (Blast_tac 1));
-result();
-
-
-
-(**** Inductive proofs about tls ****)
-
-
-(*Induction for regularity theorems.  If induction formula has the form
-   X \\<notin> analz (spies evs) --> ... then it shortens the proof by discarding
-   needless information about analz (insert X (spies evs))  *)
-fun parts_induct_tac i = 
-    etac tls.induct i
-    THEN REPEAT (FIRSTGOAL analz_mono_contra_tac)
-    THEN Force_tac i THEN
-    ALLGOALS Asm_simp_tac;
-
-
-(** Theorems of the form X \\<notin> parts (spies evs) imply that NOBODY
-    sends messages containing X! **)
-
-(*Spy never sees another agent's private key! (unless it's bad at start)*)
-Goal "evs \\<in> tls ==> (Key (priK A) \\<in> parts (spies evs)) = (A \\<in> bad)";
-by (parts_induct_tac 1);
-by (Blast_tac 1);
-qed "Spy_see_priK";
-Addsimps [Spy_see_priK];
-
-Goal "evs \\<in> tls ==> (Key (priK A) \\<in> analz (spies evs)) = (A \\<in> bad)";
-by Auto_tac;
-qed "Spy_analz_priK";
-Addsimps [Spy_analz_priK];
-
-AddSDs [Spy_see_priK RSN (2, rev_iffD1), 
-	Spy_analz_priK RSN (2, rev_iffD1)];
-
-
-(*This lemma says that no false certificates exist.  One might extend the
-  model to include bogus certificates for the agents, but there seems
-  little point in doing so: the loss of their private keys is a worse
-  breach of security.*)
-Goalw [certificate_def]
-    "[| certificate B KB \\<in> parts (spies evs);  evs \\<in> tls |] ==> pubK B = KB";
-by (etac rev_mp 1);
-by (parts_induct_tac 1);
-by (Blast_tac 1);
-qed "certificate_valid";
-
-
-(*Replace key KB in ClientKeyExch by (pubK B) *)
-val ClientKeyExch_tac = 
-    forward_tac [Says_imp_spies RS parts.Inj RS certificate_valid]
-    THEN' assume_tac
-    THEN' hyp_subst_tac;
-
-fun analz_induct_tac i = 
-    etac tls.induct i   THEN
-    ClientKeyExch_tac  (i+6)  THEN	(*ClientKeyExch*)
-    ALLGOALS (asm_simp_tac (simpset() addsimps split_ifs @ pushes))  THEN
-    (*Remove instances of pubK B:  the Spy already knows all public keys.
-      Combining the two simplifier calls makes them run extremely slowly.*)
-    ALLGOALS (asm_simp_tac (simpset() addsimps [image_eq_UN, insert_absorb]));
-
-
-(*** Properties of items found in Notes ***)
-
-Goal "[| Notes A {|Agent B, X|} \\<in> set evs;  evs \\<in> tls |]  \
-\     ==> Crypt (pubK B) X \\<in> parts (spies evs)";
-by (etac rev_mp 1);
-by (analz_induct_tac 1);
-by (blast_tac (claset() addIs [parts_insertI]) 1);
-qed "Notes_Crypt_parts_spies";
-
-(*C may be either A or B*)
-Goal "[| Notes C {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \\<in> set evs; \
-\        evs \\<in> tls |]    \
-\     ==> Crypt (pubK B) (Nonce PMS) \\<in> parts (spies evs)";
-by (etac rev_mp 1);
-by (parts_induct_tac 1);
-by (ALLGOALS Clarify_tac);
-(*Fake*)
-by (blast_tac (claset() addIs [parts_insertI]) 1);
-(*Client, Server Accept*)
-by (REPEAT (blast_tac (claset() addSDs [Notes_Crypt_parts_spies]) 1));
-qed "Notes_master_imp_Crypt_PMS";
-
-(*Compared with the theorem above, both premise and conclusion are stronger*)
-Goal "[| Notes A {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \\<in> set evs;\
-\        evs \\<in> tls |]    \
-\     ==> Notes A {|Agent B, Nonce PMS|} \\<in> set evs";
-by (etac rev_mp 1);
-by (parts_induct_tac 1);
-(*ServerAccepts*)
-by (Fast_tac 1);
-qed "Notes_master_imp_Notes_PMS";
-
-
-(*** Protocol goal: if B receives CertVerify, then A sent it ***)
-
-(*B can check A's signature if he has received A's certificate.*)
-Goal "[| X \\<in> parts (spies evs);                          \
-\        X = Crypt (priK A) (Hash{|nb, Agent B, pms|});  \
-\        evs \\<in> tls;  A \\<notin> bad |]                         \
-\     ==> Says A B X \\<in> set evs";
-by (etac rev_mp 1);
-by (hyp_subst_tac 1);
-by (parts_induct_tac 1);
-by (Blast_tac 1);
-val lemma = result();
-
-(*Final version: B checks X using the distributed KA instead of priK A*)
-Goal "[| X \\<in> parts (spies evs);                            \
-\        X = Crypt (invKey KA) (Hash{|nb, Agent B, pms|}); \
-\        certificate A KA \\<in> parts (spies evs);             \
-\        evs \\<in> tls;  A \\<notin> bad |]                           \
-\     ==> Says A B X \\<in> set evs";
-by (blast_tac (claset() addSDs [certificate_valid] addSIs [lemma]) 1);
-qed "TrustCertVerify";
-
-
-(*If CertVerify is present then A has chosen PMS.*)
-Goal "[| Crypt (priK A) (Hash{|nb, Agent B, Nonce PMS|}) \
-\          \\<in> parts (spies evs);                          \
-\        evs \\<in> tls;  A \\<notin> bad |]                         \
-\     ==> Notes A {|Agent B, Nonce PMS|} \\<in> set evs";
-by (etac rev_mp 1);
-by (parts_induct_tac 1);
-by (Blast_tac 1);
-val lemma = result();
-
-(*Final version using the distributed KA instead of priK A*)
-Goal "[| Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|}) \
-\          \\<in> parts (spies evs);                             \
-\        certificate A KA \\<in> parts (spies evs);              \
-\        evs \\<in> tls;  A \\<notin> bad |]                            \
-\     ==> Notes A {|Agent B, Nonce PMS|} \\<in> set evs";
-by (blast_tac (claset() addSDs [certificate_valid] addSIs [lemma]) 1);
-qed "UseCertVerify";
-
-
-Goal "evs \\<in> tls ==> Notes A {|Agent B, Nonce (PRF x)|} \\<notin> set evs";
-by (parts_induct_tac 1);
-(*ClientKeyExch: PMS is assumed to differ from any PRF.*)
-by (Blast_tac 1);
-qed "no_Notes_A_PRF";
-Addsimps [no_Notes_A_PRF];
-
-
-Goal "[| Nonce (PRF (PMS,NA,NB)) \\<in> parts (spies evs);  evs \\<in> tls |]  \
-\     ==> Nonce PMS \\<in> parts (spies evs)";
-by (etac rev_mp 1);
-by (parts_induct_tac 1);
-(*Easy, e.g. by freshness*)
-by (REPEAT (blast_tac (claset() addDs [Notes_Crypt_parts_spies]) 2));
-(*Fake*)
-by (blast_tac (claset() addIs [parts_insertI]) 1);
-qed "MS_imp_PMS";
-AddSDs [MS_imp_PMS];
-
-
-
-(*** Unicity results for PMS, the pre-master-secret ***)
-
-(*PMS determines B.*)
-Goal "[| Crypt(pubK B)  (Nonce PMS) \\<in> parts (spies evs); \
-\        Crypt(pubK B') (Nonce PMS) \\<in> parts (spies evs); \
-\        Nonce PMS \\<notin> analz (spies evs);                 \
-\        evs \\<in> tls |]                                          \
-\     ==> B=B'";
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (parts_induct_tac 1);
-(*Fake, ClientKeyExch*)
-by (ALLGOALS Blast_tac);
-qed "Crypt_unique_PMS";
-
-
-(** It is frustrating that we need two versions of the unicity results.
-    But Notes A {|Agent B, Nonce PMS|} determines both A and B.  Sometimes
-    we have only the weaker assertion Crypt(pubK B) (Nonce PMS), which 
-    determines B alone, and only if PMS is secret.
-**)
-
-(*In A's internal Note, PMS determines A and B.*)
-Goal "[| Notes A  {|Agent B,  Nonce PMS|} \\<in> set evs;  \
-\        Notes A' {|Agent B', Nonce PMS|} \\<in> set evs;  \
-\        evs \\<in> tls |]                               \
-\     ==> A=A' & B=B'";
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (parts_induct_tac 1);
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [Notes_Crypt_parts_spies]) 1);
-qed "Notes_unique_PMS";
-
-
-(**** Secrecy Theorems ****)
-
-(*Key compromise lemma needed to prove analz_image_keys.
-  No collection of keys can help the spy get new private keys.*)
-Goal "evs \\<in> tls                                      \
-\     ==> \\<forall>KK. (Key(priK B) \\<in> analz (Key`KK Un (spies evs))) = \
-\         (priK B \\<in> KK | B \\<in> bad)";
-by (etac tls.induct 1);
-by (ALLGOALS
-    (asm_simp_tac (analz_image_keys_ss
-		   addsimps certificate_def::keys_distinct)));
-(*Fake*) 
-by (spy_analz_tac 1);
-qed_spec_mp "analz_image_priK";
-
-
-(*slightly speeds up the big simplification below*)
-Goal "KK <= range sessionK ==> priK B \\<notin> KK";
-by (Blast_tac 1);
-val range_sessionkeys_not_priK = result();
-
-(*Lemma for the trivial direction of the if-and-only-if*)
-Goal "(X \\<in> analz (G Un H)) --> (X \\<in> analz H)  ==> \
-\     (X \\<in> analz (G Un H))  =  (X \\<in> analz H)";
-by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1);
-val analz_image_keys_lemma = result();
-
-(** Strangely, the following version doesn't work:
-\ \\<forall>Z. (Nonce N \\<in> analz (Key`(sessionK`Z) Un (spies evs))) = \
-\        (Nonce N \\<in> analz (spies evs))";
-**)
-
-Goal "evs \\<in> tls ==>                                    \
-\ \\<forall>KK. KK <= range sessionK -->                     \
-\         (Nonce N \\<in> analz (Key`KK Un (spies evs))) = \
-\         (Nonce N \\<in> analz (spies evs))";
-by (etac tls.induct 1);
-by (ClientKeyExch_tac 7);
-by (REPEAT_FIRST (resolve_tac [allI, impI]));
-by (REPEAT_FIRST (rtac analz_image_keys_lemma));
-by (ALLGOALS    (*4.5 seconds*)
-    (asm_simp_tac (analz_image_keys_ss 
-		   addsimps split_ifs @ pushes @
-		            [range_sessionkeys_not_priK, 
-			     analz_image_priK, certificate_def])));
-by (ALLGOALS (asm_simp_tac (simpset() addsimps [insert_absorb])));
-(*Fake*) 
-by (spy_analz_tac 1);
-qed_spec_mp "analz_image_keys";
-
-(*Knowing some session keys is no help in getting new nonces*)
-Goal "evs \\<in> tls ==>          \
-\     Nonce N \\<in> analz (insert (Key (sessionK z)) (spies evs)) =  \
-\     (Nonce N \\<in> analz (spies evs))";
-by (asm_simp_tac (analz_image_keys_ss addsimps [analz_image_keys]) 1);
-qed "analz_insert_key";
-Addsimps [analz_insert_key];
-
-
-(*** Protocol goal: serverK(Na,Nb,M) and clientK(Na,Nb,M) remain secure ***)
-
-(** Some lemmas about session keys, comprising clientK and serverK **)
-
-
-(*Lemma: session keys are never used if PMS is fresh.  
-  Nonces don't have to agree, allowing session resumption.
-  Converse doesn't hold; revealing PMS doesn't force the keys to be sent.
-  THEY ARE NOT SUITABLE AS SAFE ELIM RULES.*)
-Goal "[| Nonce PMS \\<notin> parts (spies evs);  \
-\        K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role);  \
-\        evs \\<in> tls |]             \
-\  ==> Key K \\<notin> parts (spies evs) & (\\<forall>Y. Crypt K Y \\<notin> parts (spies evs))";
-by (etac rev_mp 1);
-by (hyp_subst_tac 1);
-by (analz_induct_tac 1);
-(*SpyKeys*)
-by (Blast_tac 2);
-(*Fake*)
-by (blast_tac (claset() addIs [parts_insertI]) 1);
-(** LEVEL 6 **)
-(*Oops*)
-by (REPEAT 
-    (force_tac (claset() addSDs [Notes_Crypt_parts_spies, 
-				 Notes_master_imp_Crypt_PMS],
-		simpset()) 1));
-val lemma = result();
-
-Goal "[| Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) \\<in> parts (spies evs); \
-\        evs \\<in> tls |]             \
-\     ==> Nonce PMS \\<in> parts (spies evs)";
-by (blast_tac (claset() addDs [lemma]) 1);
-qed "PMS_sessionK_not_spied";
-
-Goal "[| Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) Y  \
-\          \\<in> parts (spies evs);  evs \\<in> tls |]             \
-\     ==> Nonce PMS \\<in> parts (spies evs)";
-by (blast_tac (claset() addDs [lemma]) 1);
-qed "PMS_Crypt_sessionK_not_spied";
-
-(*Write keys are never sent if M (MASTER SECRET) is secure.  
-  Converse fails; betraying M doesn't force the keys to be sent!
-  The strong Oops condition can be weakened later by unicity reasoning, 
-  with some effort.  
-  NO LONGER USED: see clientK_not_spied and serverK_not_spied*)
-Goal "[| \\<forall>A. Says A Spy (Key (sessionK((NA,NB,M),role))) \\<notin> set evs; \
-\        Nonce M \\<notin> analz (spies evs);  evs \\<in> tls |]   \
-\     ==> Key (sessionK((NA,NB,M),role)) \\<notin> parts (spies evs)";
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (analz_induct_tac 1);        (*5 seconds*)
-(*SpyKeys*)
-by (blast_tac (claset() addDs [Says_imp_spies RS analz.Inj]) 2);
-(*Fake*) 
-by (spy_analz_tac 1);
-qed "sessionK_not_spied";
-
-
-(*If A sends ClientKeyExch to an honest B, then the PMS will stay secret.*)
-Goal "[| evs \\<in> tls;  A \\<notin> bad;  B \\<notin> bad |]           \
-\     ==> Notes A {|Agent B, Nonce PMS|} \\<in> set evs  -->   \
-\         Nonce PMS \\<notin> analz (spies evs)";
-by (analz_induct_tac 1);   (*4 seconds*)
-(*ClientAccepts and ServerAccepts: because PMS \\<notin> range PRF*)
-by (REPEAT (Force_tac 6));
-(*ClientHello, ServerHello, ClientKeyExch, ServerResume: 
-  mostly freshness reasoning*)
-by (REPEAT (blast_tac (claset() addSDs [parts.Body]
-				addDs  [Notes_Crypt_parts_spies,
-					Says_imp_spies RS analz.Inj]) 3));
-(*SpyKeys*)
-by (Force_tac 2);
-(*Fake*)
-by (spy_analz_tac 1);
-bind_thm ("Spy_not_see_PMS", result() RSN (2, rev_mp));
-
-
-(*If A sends ClientKeyExch to an honest B, then the MASTER SECRET
-  will stay secret.*)
-Goal "[| evs \\<in> tls;  A \\<notin> bad;  B \\<notin> bad |]           \
-\     ==> Notes A {|Agent B, Nonce PMS|} \\<in> set evs  -->   \
-\         Nonce (PRF(PMS,NA,NB)) \\<notin> analz (spies evs)";
-by (analz_induct_tac 1);   (*4 seconds*)
-(*ClientAccepts and ServerAccepts: because PMS was already visible*)
-by (REPEAT (blast_tac (claset() addDs [Spy_not_see_PMS, 
-				       Says_imp_spies RS analz.Inj,
-				       Notes_imp_knows_Spy RS analz.Inj]) 6));
-(*ClientHello*)
-by (Blast_tac 3);
-(*SpyKeys: by secrecy of the PMS, Spy cannot make the MS*)
-by (blast_tac (claset() addSDs [Spy_not_see_PMS, 
-				Says_imp_spies RS analz.Inj]) 2);
-(*Fake*)
-by (spy_analz_tac 1);
-(*ServerHello and ClientKeyExch: mostly freshness reasoning*)
-by (REPEAT (blast_tac (claset() addSDs [parts.Body]
-				addDs  [Notes_Crypt_parts_spies,
-					Says_imp_spies RS analz.Inj]) 1));
-bind_thm ("Spy_not_see_MS", result() RSN (2, rev_mp));
-
-
-(*** Weakening the Oops conditions for leakage of clientK ***)
-
-(*If A created PMS then nobody else (except the Spy in replays) 
-  would send a message using a clientK generated from that PMS.*)
-Goal "[| Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) \\<in> set evs;  \
-\        Notes A {|Agent B, Nonce PMS|} \\<in> set evs;   \
-\        evs \\<in> tls;  A' \\<noteq> Spy |]                \
-\     ==> A = A'";
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (analz_induct_tac 1); 
-by (ALLGOALS Clarify_tac);
-(*ClientFinished, ClientResume: by unicity of PMS*)
-by (REPEAT 
-    (blast_tac (claset() addSDs [Notes_master_imp_Notes_PMS]
-     	 	         addIs  [Notes_unique_PMS RS conjunct1]) 2));
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied,
-				Says_imp_spies RS parts.Inj]) 1);
-qed "Says_clientK_unique";
-
-
-(*If A created PMS and has not leaked her clientK to the Spy, 
-  then it is completely secure: not even in parts!*)
-Goal "[| Notes A {|Agent B, Nonce PMS|} \\<in> set evs;  \
-\        Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) \\<notin> set evs;  \
-\        A \\<notin> bad;  B \\<notin> bad; \
-\        evs \\<in> tls |]   \
-\     ==> Key (clientK(Na,Nb,PRF(PMS,NA,NB))) \\<notin> parts (spies evs)";
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (analz_induct_tac 1);        (*4 seconds*)
-(*Oops*)
-by (blast_tac (claset() addIs [Says_clientK_unique]) 4);
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [PMS_sessionK_not_spied]) 3);
-(*SpyKeys*)
-by (blast_tac (claset() addSEs [Spy_not_see_MS RSN (2,rev_notE)]) 2);
-(*Fake*) 
-by (spy_analz_tac 1);
-qed "clientK_not_spied";
-
-
-(*** Weakening the Oops conditions for leakage of serverK ***)
-
-(*If A created PMS for B, then nobody other than B or the Spy would
-  send a message using a serverK generated from that PMS.*)
-Goal "[| Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) \\<in> set evs;  \
-\        Notes A {|Agent B, Nonce PMS|} \\<in> set evs;  \
-\        evs \\<in> tls;  A \\<notin> bad;  B \\<notin> bad;  B' \\<noteq> Spy |]                \
-\     ==> B = B'";
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (analz_induct_tac 1);
-by (ALLGOALS Clarify_tac);
-(*ServerResume, ServerFinished: by unicity of PMS*)
-by (REPEAT
-    (blast_tac (claset() addSDs [Notes_master_imp_Crypt_PMS]
-			 addDs  [Spy_not_see_PMS, 
-				 Notes_Crypt_parts_spies,
-				 Crypt_unique_PMS]) 2));
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied,
-				Says_imp_spies RS parts.Inj]) 1);
-qed "Says_serverK_unique";
-
-(*If A created PMS for B, and B has not leaked his serverK to the Spy, 
-  then it is completely secure: not even in parts!*)
-Goal "[| Notes A {|Agent B, Nonce PMS|} \\<in> set evs;                   \
-\        Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) \\<notin> set evs; \
-\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> tls |]                          \
-\     ==> Key (serverK(Na,Nb,PRF(PMS,NA,NB))) \\<notin> parts (spies evs)";
-by (etac rev_mp 1);
-by (etac rev_mp 1);
-by (analz_induct_tac 1);
-(*Oops*)
-by (blast_tac (claset() addIs [Says_serverK_unique]) 4);
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [PMS_sessionK_not_spied]) 3);
-(*SpyKeys*)
-by (blast_tac (claset() addSEs [Spy_not_see_MS RSN (2,rev_notE)]) 2);
-(*Fake*) 
-by (spy_analz_tac 1);
-qed "serverK_not_spied";
-
-
-(*** Protocol goals: if A receives ServerFinished, then B is present 
-     and has used the quoted values PA, PB, etc.  Note that it is up to A
-     to compare PA with what she originally sent.
-***)
-
-(*The mention of her name (A) in X assures A that B knows who she is.*)
-Goal "[| X = Crypt (serverK(Na,Nb,M))                  \
-\              (Hash{|Number SID, Nonce M,             \
-\                     Nonce Na, Number PA, Agent A,    \
-\                     Nonce Nb, Number PB, Agent B|}); \
-\        M = PRF(PMS,NA,NB);                           \
-\        evs \\<in> tls;  A \\<notin> bad;  B \\<notin> bad |]            \
-\     ==> Says B Spy (Key(serverK(Na,Nb,M))) \\<notin> set evs --> \
-\         Notes A {|Agent B, Nonce PMS|} \\<in> set evs --> \
-\         X \\<in> parts (spies evs) --> Says B A X \\<in> set evs";
-by (hyp_subst_tac 1);
-by (analz_induct_tac 1);        (*7 seconds*)
-by (ALLGOALS Clarify_tac);
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied]) 2);
-(*Fake: the Spy doesn't have the critical session key!*)
-by (blast_tac (claset() addEs [serverK_not_spied RSN (2,rev_notE)]) 1);
-qed_spec_mp "TrustServerFinished";
-
-(*This version refers not to ServerFinished but to any message from B.
-  We don't assume B has received CertVerify, and an intruder could
-  have changed A's identity in all other messages, so we can't be sure
-  that B sends his message to A.  If CLIENT KEY EXCHANGE were augmented
-  to bind A's identity with PMS, then we could replace A' by A below.*)
-Goal "[| M = PRF(PMS,NA,NB);  evs \\<in> tls;  A \\<notin> bad;  B \\<notin> bad |]     \
-\     ==> Says B Spy (Key(serverK(Na,Nb,M))) \\<notin> set evs --> \
-\         Notes A {|Agent B, Nonce PMS|} \\<in> set evs -->              \
-\         Crypt (serverK(Na,Nb,M)) Y \\<in> parts (spies evs)  -->  \
-\         (\\<exists>A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) \\<in> set evs)";
-by (hyp_subst_tac 1);
-by (analz_induct_tac 1);	(*6 seconds*)
-by (ALLGOALS (asm_simp_tac (simpset() addsimps [ex_disj_distrib])));
-by (ALLGOALS Clarify_tac);
-(*ServerResume, ServerFinished: by unicity of PMS*)
-by (REPEAT
-    (blast_tac (claset() addSDs [Notes_master_imp_Crypt_PMS]
-			 addDs  [Spy_not_see_PMS, 
-				 Notes_Crypt_parts_spies,
-				 Crypt_unique_PMS]) 3));
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied]) 2);
-(*Fake: the Spy doesn't have the critical session key!*)
-by (blast_tac (claset() addEs [serverK_not_spied RSN (2,rev_notE)]) 1);
-qed_spec_mp "TrustServerMsg";
-
-
-(*** Protocol goal: if B receives any message encrypted with clientK
-     then A has sent it, ASSUMING that A chose PMS.  Authentication is
-     assumed here; B cannot verify it.  But if the message is
-     ClientFinished, then B can then check the quoted values PA, PB, etc.
-***)
-
-Goal "[| M = PRF(PMS,NA,NB);  evs \\<in> tls;  A \\<notin> bad;  B \\<notin> bad |] \
-\     ==> Says A Spy (Key(clientK(Na,Nb,M))) \\<notin> set evs --> \
-\         Notes A {|Agent B, Nonce PMS|} \\<in> set evs -->               \
-\         Crypt (clientK(Na,Nb,M)) Y \\<in> parts (spies evs) -->         \
-\         Says A B (Crypt (clientK(Na,Nb,M)) Y) \\<in> set evs";
-by (hyp_subst_tac 1);
-by (analz_induct_tac 1);	(*6 seconds*)
-by (ALLGOALS Clarify_tac);
-(*ClientFinished, ClientResume: by unicity of PMS*)
-by (REPEAT (blast_tac (claset() delrules [conjI]
-		                addSDs [Notes_master_imp_Notes_PMS]
-	 	                addDs  [Notes_unique_PMS]) 3));
-(*ClientKeyExch*)
-by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied]) 2);
-(*Fake: the Spy doesn't have the critical session key!*)
-by (blast_tac (claset() addEs [clientK_not_spied RSN (2,rev_notE)]) 1);
-qed_spec_mp "TrustClientMsg";
-
-
-
-(*** Protocol goal: if B receives ClientFinished, and if B is able to
-     check a CertVerify from A, then A has used the quoted
-     values PA, PB, etc.  Even this one requires A to be uncompromised.
- ***)
-Goal "[| M = PRF(PMS,NA,NB);                           \
-\        Says A Spy (Key(clientK(Na,Nb,M))) \\<notin> set evs;\
-\        Says A' B (Crypt (clientK(Na,Nb,M)) Y) \\<in> set evs; \
-\        certificate A KA \\<in> parts (spies evs);       \
-\        Says A'' B (Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|}))\
-\          \\<in> set evs;                                                  \
-\        evs \\<in> tls;  A \\<notin> bad;  B \\<notin> bad |]                             \
-\     ==> Says A B (Crypt (clientK(Na,Nb,M)) Y) \\<in> set evs";
-by (blast_tac (claset() addSIs [TrustClientMsg, UseCertVerify]
-                        addDs  [Says_imp_spies RS parts.Inj]) 1);
-qed "AuthClientFinished";
-
-(*22/9/97: loads in 622s, which is 10 minutes 22 seconds*)
-(*24/9/97: loads in 672s, which is 11 minutes 12 seconds [stronger theorems]*)
-(*29/9/97: loads in 481s, after removing Certificate from ClientKeyExch*)
-(*30/9/97: loads in 476s, after removing unused theorems*)
-(*30/9/97: loads in 448s, after fixing ServerResume*)
-
-(*08/9/97: loads in 189s (pike), after much reorganization, 
-           back to 621s on albatross?*)
-
-(*10/2/99: loads in 139s (pike)
-           down to 433s on albatross*)
--- a/src/HOL/Auth/TLS.thy	Mon May 07 19:19:41 2001 +0200
+++ b/src/HOL/Auth/TLS.thy	Tue May 08 15:56:57 2001 +0200
@@ -39,7 +39,7 @@
 	Notes A {|Agent B, Nonce PMS|}.
 *)
 
-TLS = Public + 
+theory TLS = Public:
 
 constdefs
   certificate      :: "[agent,key] => msg"
@@ -58,196 +58,853 @@
   sessionK :: "(nat*nat*nat) * role => key"
 
 syntax
-    clientK, serverK :: "nat*nat*nat => key"
+    clientK :: "nat*nat*nat => key"
+    serverK :: "nat*nat*nat => key"
 
 translations
   "clientK X" == "sessionK(X, ClientRole)"
   "serverK X" == "sessionK(X, ServerRole)"
 
-rules
+axioms
   (*the pseudo-random function is collision-free*)
-  inj_PRF       "inj PRF"	
+  inj_PRF:       "inj PRF"
 
   (*sessionK is collision-free; also, no clientK clashes with any serverK.*)
-  inj_sessionK  "inj sessionK"	
+  inj_sessionK:  "inj sessionK"	
 
   (*sessionK makes symmetric keys*)
-  isSym_sessionK "sessionK nonces \\<in> symKeys"
+  isSym_sessionK: "sessionK nonces \<in> symKeys"
 
 
-consts    tls :: event list set
+consts    tls :: "event list set"
 inductive tls
-  intrs 
-    Nil  (*Initial trace is empty*)
-         "[]: tls"
+  intros
+   Nil:  (*Initial trace is empty*)
+         "[] \<in> tls"
 
-    Fake (*The spy, an active attacker, MAY say anything he CAN say.*)
-         "[| evsf \\<in> tls;  X \\<in> synth (analz (spies evsf)) |]
-          ==> Says Spy B X # evsf \\<in> tls"
+   Fake: (*The spy, an active attacker, MAY say anything he CAN say.*)
+         "[| evsf \<in> tls;  X \<in> synth (analz (spies evsf)) |]
+          ==> Says Spy B X # evsf \<in> tls"
 
-    SpyKeys (*The spy may apply PRF & sessionK to available nonces*)
-         "[| evsSK \\<in> tls;
+   SpyKeys: (*The spy may apply PRF & sessionK to available nonces*)
+         "[| evsSK \<in> tls;
 	     {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
           ==> Notes Spy {| Nonce (PRF(M,NA,NB)),
-			   Key (sessionK((NA,NB,M),role)) |} # evsSK \\<in> tls"
+			   Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
 
-    ClientHello
+   ClientHello:
 	 (*(7.4.1.2)
 	   PA represents CLIENT_VERSION, CIPHER_SUITES and COMPRESSION_METHODS.
 	   It is uninterpreted but will be confirmed in the FINISHED messages.
 	   NA is CLIENT RANDOM, while SID is SESSION_ID.
            UNIX TIME is omitted because the protocol doesn't use it.
-           May assume NA \\<notin> range PRF because CLIENT RANDOM is 28 bytes
+           May assume NA \<notin> range PRF because CLIENT RANDOM is 28 bytes
 	   while MASTER SECRET is 48 bytes*)
-         "[| evsCH \\<in> tls;  Nonce NA \\<notin> used evsCH;  NA \\<notin> range PRF |]
+         "[| evsCH \<in> tls;  Nonce NA \<notin> used evsCH;  NA \<notin> range PRF |]
           ==> Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
-	        # evsCH  \\<in>  tls"
+	        # evsCH  \<in>  tls"
 
-    ServerHello
+   ServerHello:
          (*7.4.1.3 of the TLS Internet-Draft
 	   PB represents CLIENT_VERSION, CIPHER_SUITE and COMPRESSION_METHOD.
            SERVER CERTIFICATE (7.4.2) is always present.
            CERTIFICATE_REQUEST (7.4.4) is implied.*)
-         "[| evsSH \\<in> tls;  Nonce NB \\<notin> used evsSH;  NB \\<notin> range PRF;
+         "[| evsSH \<in> tls;  Nonce NB \<notin> used evsSH;  NB \<notin> range PRF;
              Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \\<in> set evsSH |]
-          ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH  \\<in>  tls"
+	       \<in> set evsSH |]
+          ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH  \<in>  tls"
 
-    Certificate
+   Certificate:
          (*SERVER (7.4.2) or CLIENT (7.4.6) CERTIFICATE.*)
-         "evsC \\<in> tls ==> Says B A (certificate B (pubK B)) # evsC  \\<in>  tls"
+         "evsC \<in> tls ==> Says B A (certificate B (pubK B)) # evsC  \<in>  tls"
 
-    ClientKeyExch
+   ClientKeyExch:
          (*CLIENT KEY EXCHANGE (7.4.7).
            The client, A, chooses PMS, the PREMASTER SECRET.
            She encrypts PMS using the supplied KB, which ought to be pubK B.
-           We assume PMS \\<notin> range PRF because a clash betweem the PMS
+           We assume PMS \<notin> range PRF because a clash betweem the PMS
            and another MASTER SECRET is highly unlikely (even though
 	   both items have the same length, 48 bytes).
            The Note event records in the trace that she knows PMS
                (see REMARK at top). *)
-         "[| evsCX \\<in> tls;  Nonce PMS \\<notin> used evsCX;  PMS \\<notin> range PRF;
-             Says B' A (certificate B KB) \\<in> set evsCX |]
+         "[| evsCX \<in> tls;  Nonce PMS \<notin> used evsCX;  PMS \<notin> range PRF;
+             Says B' A (certificate B KB) \<in> set evsCX |]
           ==> Says A B (Crypt KB (Nonce PMS))
 	      # Notes A {|Agent B, Nonce PMS|}
-	      # evsCX  \\<in>  tls"
+	      # evsCX  \<in>  tls"
 
-    CertVerify
+   CertVerify:
 	(*The optional Certificate Verify (7.4.8) message contains the
           specific components listed in the security analysis, F.1.1.2.
           It adds the pre-master-secret, which is also essential!
           Checking the signature, which is the only use of A's certificate,
           assures B of A's presence*)
-         "[| evsCV \\<in> tls;  
-             Says B' A {|Nonce NB, Number SID, Number PB|} \\<in> set evsCV;
-	     Notes A {|Agent B, Nonce PMS|} \\<in> set evsCV |]
+         "[| evsCV \<in> tls;
+             Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCV;
+	     Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
           ==> Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}))
-              # evsCV  \\<in>  tls"
+              # evsCV  \<in>  tls"
 
 	(*Finally come the FINISHED messages (7.4.8), confirming PA and PB
           among other things.  The master-secret is PRF(PMS,NA,NB).
-          Either party may sent its message first.*)
+          Either party may send its message first.*)
 
-    ClientFinished
-        (*The occurrence of Notes A {|Agent B, Nonce PMS|} stops the 
+   ClientFinished:
+        (*The occurrence of Notes A {|Agent B, Nonce PMS|} stops the
           rule's applying when the Spy has satisfied the "Says A B" by
           repaying messages sent by the true client; in that case, the
           Spy does not know PMS and could not send ClientFinished.  One
-          could simply put A\\<noteq>Spy into the rule, but one should not
+          could simply put A\<noteq>Spy into the rule, but one should not
           expect the spy to be well-behaved.*)
-         "[| evsCF \\<in> tls;  
+         "[| evsCF \<in> tls;
 	     Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \\<in> set evsCF;
-             Says B' A {|Nonce NB, Number SID, Number PB|} \\<in> set evsCF;
-             Notes A {|Agent B, Nonce PMS|} \\<in> set evsCF;
+	       \<in> set evsCF;
+             Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCF;
+             Notes A {|Agent B, Nonce PMS|} \<in> set evsCF;
 	     M = PRF(PMS,NA,NB) |]
           ==> Says A B (Crypt (clientK(NA,NB,M))
 			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A, 
+			       Nonce NA, Number PA, Agent A,
 			       Nonce NB, Number PB, Agent B|}))
-              # evsCF  \\<in>  tls"
+              # evsCF  \<in>  tls"
 
-    ServerFinished
+   ServerFinished:
 	(*Keeping A' and A'' distinct means B cannot even check that the
           two messages originate from the same source. *)
-         "[| evsSF \\<in> tls;
+         "[| evsSF \<in> tls;
 	     Says A' B  {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \\<in> set evsSF;
-	     Says B  A  {|Nonce NB, Number SID, Number PB|} \\<in> set evsSF;
-	     Says A'' B (Crypt (pubK B) (Nonce PMS)) \\<in> set evsSF;
+	       \<in> set evsSF;
+	     Says B  A  {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
+	     Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
 	     M = PRF(PMS,NA,NB) |]
           ==> Says B A (Crypt (serverK(NA,NB,M))
 			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A, 
+			       Nonce NA, Number PA, Agent A,
 			       Nonce NB, Number PB, Agent B|}))
-              # evsSF  \\<in>  tls"
+              # evsSF  \<in>  tls"
 
-    ClientAccepts
+   ClientAccepts:
 	(*Having transmitted ClientFinished and received an identical
           message encrypted with serverK, the client stores the parameters
           needed to resume this session.  The "Notes A ..." premise is
           used to prove Notes_master_imp_Crypt_PMS.*)
-         "[| evsCA \\<in> tls;
-	     Notes A {|Agent B, Nonce PMS|} \\<in> set evsCA;
-	     M = PRF(PMS,NA,NB);  
+         "[| evsCA \<in> tls;
+	     Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
+	     M = PRF(PMS,NA,NB);
 	     X = Hash{|Number SID, Nonce M,
-	               Nonce NA, Number PA, Agent A, 
+	               Nonce NA, Number PA, Agent A,
 		       Nonce NB, Number PB, Agent B|};
-             Says A  B (Crypt (clientK(NA,NB,M)) X) \\<in> set evsCA;
-             Says B' A (Crypt (serverK(NA,NB,M)) X) \\<in> set evsCA |]
-          ==> 
-             Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA  \\<in>  tls"
+             Says A  B (Crypt (clientK(NA,NB,M)) X) \<in> set evsCA;
+             Says B' A (Crypt (serverK(NA,NB,M)) X) \<in> set evsCA |]
+          ==>
+             Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA  \<in>  tls"
 
-    ServerAccepts
+   ServerAccepts:
 	(*Having transmitted ServerFinished and received an identical
           message encrypted with clientK, the server stores the parameters
           needed to resume this session.  The "Says A'' B ..." premise is
           used to prove Notes_master_imp_Crypt_PMS.*)
-         "[| evsSA \\<in> tls;
-	     A \\<noteq> B;
-             Says A'' B (Crypt (pubK B) (Nonce PMS)) \\<in> set evsSA;
-	     M = PRF(PMS,NA,NB);  
+         "[| evsSA \<in> tls;
+	     A \<noteq> B;
+             Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSA;
+	     M = PRF(PMS,NA,NB);
 	     X = Hash{|Number SID, Nonce M,
-	               Nonce NA, Number PA, Agent A, 
+	               Nonce NA, Number PA, Agent A,
 		       Nonce NB, Number PB, Agent B|};
-             Says B  A (Crypt (serverK(NA,NB,M)) X) \\<in> set evsSA;
-             Says A' B (Crypt (clientK(NA,NB,M)) X) \\<in> set evsSA |]
-          ==> 
-             Notes B {|Number SID, Agent A, Agent B, Nonce M|} # evsSA  \\<in>  tls"
+             Says B  A (Crypt (serverK(NA,NB,M)) X) \<in> set evsSA;
+             Says A' B (Crypt (clientK(NA,NB,M)) X) \<in> set evsSA |]
+          ==>
+             Notes B {|Number SID, Agent A, Agent B, Nonce M|} # evsSA  \<in>  tls"
 
-    ClientResume
+   ClientResume:
          (*If A recalls the SESSION_ID, then she sends a FINISHED message
            using the new nonces and stored MASTER SECRET.*)
-         "[| evsCR \\<in> tls;  
+         "[| evsCR \<in> tls;
 	     Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
-             Says B' A {|Nonce NB, Number SID, Number PB|} \\<in> set evsCR;
-             Notes A {|Number SID, Agent A, Agent B, Nonce M|} \\<in> set evsCR |]
+             Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCR;
+             Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsCR |]
           ==> Says A B (Crypt (clientK(NA,NB,M))
 			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A, 
+			       Nonce NA, Number PA, Agent A,
 			       Nonce NB, Number PB, Agent B|}))
-              # evsCR  \\<in>  tls"
+              # evsCR  \<in>  tls"
 
-    ServerResume
+   ServerResume:
          (*Resumption (7.3):  If B finds the SESSION_ID then he can send
            a FINISHED message using the recovered MASTER SECRET*)
-         "[| evsSR \\<in> tls;
+         "[| evsSR \<in> tls;
 	     Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
-	     Says B  A {|Nonce NB, Number SID, Number PB|} \\<in> set evsSR;  
-             Notes B {|Number SID, Agent A, Agent B, Nonce M|} \\<in> set evsSR |]
+	     Says B  A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
+             Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsSR |]
           ==> Says B A (Crypt (serverK(NA,NB,M))
 			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A, 
+			       Nonce NA, Number PA, Agent A,
 			       Nonce NB, Number PB, Agent B|})) # evsSR
-	        \\<in>  tls"
+	        \<in>  tls"
 
-    Oops 
+   Oops:
          (*The most plausible compromise is of an old session key.  Losing
            the MASTER SECRET or PREMASTER SECRET is more serious but
-           rather unlikely.  The assumption A \\<noteq> Spy is essential: otherwise
+           rather unlikely.  The assumption A \<noteq> Spy is essential: otherwise
            the Spy could learn session keys merely by replaying messages!*)
-         "[| evso \\<in> tls;  A \\<noteq> Spy;
-	     Says A B (Crypt (sessionK((NA,NB,M),role)) X) \\<in> set evso |]
-          ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso  \\<in>  tls"
+         "[| evso \<in> tls;  A \<noteq> Spy;
+	     Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
+          ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso  \<in>  tls"
+
+(*
+Protocol goals:
+* M, serverK(NA,NB,M) and clientK(NA,NB,M) will be known only to the two
+     parties (though A is not necessarily authenticated).
+
+* B upon receiving CertVerify knows that A is present (But this
+    message is optional!)
+
+* A upon receiving ServerFinished knows that B is present
+
+* Each party who has received a FINISHED message can trust that the other
+  party agrees on all message components, including PA and PB (thus foiling
+  rollback attacks).
+*)
+
+declare Says_imp_knows_Spy [THEN analz.Inj, dest]
+declare parts.Body  [dest]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un  [dest]
+
+
+(*Automatically unfold the definition of "certificate"*)
+declare certificate_def [simp]
+
+(*Injectiveness of key-generating functions*)
+declare inj_PRF [THEN inj_eq, iff]
+declare inj_sessionK [THEN inj_eq, iff]
+declare isSym_sessionK [simp]
+
+
+(*** clientK and serverK make symmetric keys; no clashes with pubK or priK ***)
+
+lemma pubK_neq_sessionK [iff]: "pubK A \<noteq> sessionK arg"
+by (simp add: symKeys_neq_imp_neq)
+
+declare pubK_neq_sessionK [THEN not_sym, iff]
+
+lemma priK_neq_sessionK [iff]: "priK A \<noteq> sessionK arg"
+by (simp add: symKeys_neq_imp_neq)
+
+declare priK_neq_sessionK [THEN not_sym, iff]
+
+lemmas keys_distinct = pubK_neq_sessionK priK_neq_sessionK
+
+
+(**** Protocol Proofs ****)
+
+(*Possibility properties state that some traces run the protocol to the end.
+  Four paths and 12 rules are considered.*)
+
+
+(** These proofs assume that the Nonce_supply nonces
+	(which have the form  @ N. Nonce N \<notin> used evs)
+    lie outside the range of PRF.  It seems reasonable, but as it is needed
+    only for the possibility theorems, it is not taken as an axiom.
+**)
+
+
+(*Possibility property ending with ClientAccepts.*)
+lemma "[| \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B |]
+      ==> \<exists>SID M. \<exists>evs \<in> tls.
+            Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs"
+apply (intro exI bexI)
+apply (rule_tac [2] tls.Nil
+                    [THEN tls.ClientHello, THEN tls.ServerHello,
+                     THEN tls.Certificate, THEN tls.ClientKeyExch,
+                     THEN tls.ClientFinished, THEN tls.ServerFinished,
+                     THEN tls.ClientAccepts])
+apply possibility
+apply blast+
+done
+
+
+(*And one for ServerAccepts.  Either FINISHED message may come first.*)
+lemma "[| \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF; A \<noteq> B |]
+      ==> \<exists>SID NA PA NB PB M. \<exists>evs \<in> tls.
+           Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs"
+apply (intro exI bexI)
+apply (rule_tac [2] tls.Nil
+                    [THEN tls.ClientHello, THEN tls.ServerHello,
+                     THEN tls.Certificate, THEN tls.ClientKeyExch,
+                     THEN tls.ServerFinished, THEN tls.ClientFinished, 
+                     THEN tls.ServerAccepts])
+apply possibility
+apply blast+
+done
+
+
+(*Another one, for CertVerify (which is optional)*)
+lemma "[| \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;  A \<noteq> B |]
+       ==> \<exists>NB PMS. \<exists>evs \<in> tls.
+              Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|})) 
+                \<in> set evs"
+apply (intro exI bexI)
+apply (rule_tac [2] tls.Nil
+                    [THEN tls.ClientHello, THEN tls.ServerHello,
+                     THEN tls.Certificate, THEN tls.ClientKeyExch,
+                     THEN tls.CertVerify])
+apply possibility
+apply blast+
+done
+
+
+(*Another one, for session resumption (both ServerResume and ClientResume).
+  NO tls.Nil here: we refer to a previous session, not the empty trace.*)
+lemma "[| evs0 \<in> tls;
+          Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs0;
+          Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evs0;
+          \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;
+          A \<noteq> B |]
+      ==> \<exists>NA PA NB PB X. \<exists>evs \<in> tls.
+		X = Hash{|Number SID, Nonce M,
+			  Nonce NA, Number PA, Agent A,
+			  Nonce NB, Number PB, Agent B|}  &
+		Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs  &
+		Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
+apply (intro exI bexI)
+apply (rule_tac [2] tls.ClientHello
+                    [THEN tls.ServerHello,
+                     THEN tls.ServerResume, THEN tls.ClientResume])
+apply possibility
+apply blast+
+done
+
+
+(**** Inductive proofs about tls ****)
+
+
+(*Induction for regularity theorems.  If induction formula has the form
+   X \<notin> analz (spies evs) --> ... then it shortens the proof by discarding
+   needless information about analz (insert X (spies evs))  
+fun parts_induct_tac i =
+    etac tls.induct i
+    THEN REPEAT (FIRSTGOAL analz_mono_contra_tac)
+    THEN Force_tac i THEN
+    ALLGOALS Asm_simp_tac
+*)
+
+(** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY
+    sends messages containing X! **)
+
+(*Spy never sees a good agent's private key!*)
+lemma Spy_see_priK [simp]:
+     "evs \<in> tls ==> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)"
+apply (erule tls.induct, force, simp_all)
+apply blast
+done
+
+lemma Spy_analz_priK [simp]:
+     "evs \<in> tls ==> (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)"
+by auto
+
+lemma Spy_see_priK_D [dest!]:
+     "[|Key (priK A) \<in> parts (knows Spy evs);  evs \<in> tls|] ==> A \<in> bad"
+by (blast dest: Spy_see_priK)
+
+
+(*This lemma says that no false certificates exist.  One might extend the
+  model to include bogus certificates for the agents, but there seems
+  little point in doing so: the loss of their private keys is a worse
+  breach of security.*)
+lemma certificate_valid:
+    "[| certificate B KB \<in> parts (spies evs);  evs \<in> tls |] ==> KB = pubK B"
+apply (erule rev_mp)
+apply (erule tls.induct, force, simp_all)
+apply blast 
+done
+
+lemmas CX_KB_is_pubKB = Says_imp_spies [THEN parts.Inj, THEN certificate_valid]
+
+
+(*** Properties of items found in Notes ***)
+
+lemma Notes_Crypt_parts_spies:
+     "[| Notes A {|Agent B, X|} \<in> set evs;  evs \<in> tls |]
+      ==> Crypt (pubK B) X \<in> parts (spies evs)"
+apply (erule rev_mp)
+apply (erule tls.induct, 
+       frule_tac [7] CX_KB_is_pubKB, force, simp_all)
+apply (blast intro: parts_insertI)
+done
+
+(*C may be either A or B*)
+lemma Notes_master_imp_Crypt_PMS:
+     "[| Notes C {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \<in> set evs;
+         evs \<in> tls |]
+      ==> Crypt (pubK B) (Nonce PMS) \<in> parts (spies evs)"
+apply (erule rev_mp)
+apply (erule tls.induct, force, simp_all)
+(*Fake*)
+apply (blast intro: parts_insertI)
+(*Client, Server Accept*)
+apply (blast dest!: Notes_Crypt_parts_spies)+
+done
+
+(*Compared with the theorem above, both premise and conclusion are stronger*)
+lemma Notes_master_imp_Notes_PMS:
+     "[| Notes A {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \<in> set evs;
+         evs \<in> tls |]
+      ==> Notes A {|Agent B, Nonce PMS|} \<in> set evs"
+apply (erule rev_mp)
+apply (erule tls.induct, force, simp_all)
+(*ServerAccepts*)
+apply blast
+done
+
+
+(*** Protocol goal: if B receives CertVerify, then A sent it ***)
+
+(*B can check A's signature if he has received A's certificate.*)
+lemma TrustCertVerify_lemma:
+     "[| X \<in> parts (spies evs);
+         X = Crypt (priK A) (Hash{|nb, Agent B, pms|});
+         evs \<in> tls;  A \<notin> bad |]
+      ==> Says A B X \<in> set evs"
+apply (erule rev_mp, erule ssubst)
+apply (erule tls.induct, force, simp_all)
+apply blast
+done
+
+(*Final version: B checks X using the distributed KA instead of priK A*)
+lemma TrustCertVerify:
+     "[| X \<in> parts (spies evs);
+         X = Crypt (invKey KA) (Hash{|nb, Agent B, pms|});
+         certificate A KA \<in> parts (spies evs);
+         evs \<in> tls;  A \<notin> bad |]
+      ==> Says A B X \<in> set evs"
+by (blast dest!: certificate_valid intro!: TrustCertVerify_lemma)
+
+
+(*If CertVerify is present then A has chosen PMS.*)
+lemma UseCertVerify_lemma:
+     "[| Crypt (priK A) (Hash{|nb, Agent B, Nonce PMS|}) \<in> parts (spies evs);
+         evs \<in> tls;  A \<notin> bad |]
+      ==> Notes A {|Agent B, Nonce PMS|} \<in> set evs"
+apply (erule rev_mp)
+apply (erule tls.induct, force, simp_all)
+apply blast
+done
+
+(*Final version using the distributed KA instead of priK A*)
+lemma UseCertVerify:
+     "[| Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|})
+           \<in> parts (spies evs);
+         certificate A KA \<in> parts (spies evs);
+         evs \<in> tls;  A \<notin> bad |]
+      ==> Notes A {|Agent B, Nonce PMS|} \<in> set evs"
+by (blast dest!: certificate_valid intro!: UseCertVerify_lemma)
+
+
+lemma no_Notes_A_PRF [simp]:
+     "evs \<in> tls ==> Notes A {|Agent B, Nonce (PRF x)|} \<notin> set evs"
+apply (erule tls.induct, force, simp_all)
+(*ClientKeyExch: PMS is assumed to differ from any PRF.*)
+apply blast
+done
+
+
+lemma MS_imp_PMS [dest!]:
+     "[| Nonce (PRF (PMS,NA,NB)) \<in> parts (spies evs);  evs \<in> tls |]
+      ==> Nonce PMS \<in> parts (spies evs)"
+apply (erule rev_mp)
+apply (erule tls.induct, force, simp_all)
+(*Fake*)
+apply (blast intro: parts_insertI)
+(*Easy, e.g. by freshness*)
+apply (blast dest: Notes_Crypt_parts_spies)+
+done
+
+
+
+
+(*** Unicity results for PMS, the pre-master-secret ***)
+
+(*PMS determines B.*)
+lemma Crypt_unique_PMS:
+     "[| Crypt(pubK B)  (Nonce PMS) \<in> parts (spies evs);
+         Crypt(pubK B') (Nonce PMS) \<in> parts (spies evs);
+         Nonce PMS \<notin> analz (spies evs);
+         evs \<in> tls |]
+      ==> B=B'"
+apply (erule rev_mp, erule rev_mp, erule rev_mp)
+apply (erule tls.induct, analz_mono_contra, force, simp_all (no_asm_simp))
+(*Fake, ClientKeyExch*)
+apply blast+
+done
+
+
+(** It is frustrating that we need two versions of the unicity results.
+    But Notes A {|Agent B, Nonce PMS|} determines both A and B.  Sometimes
+    we have only the weaker assertion Crypt(pubK B) (Nonce PMS), which
+    determines B alone, and only if PMS is secret.
+**)
+
+(*In A's internal Note, PMS determines A and B.*)
+lemma Notes_unique_PMS:
+     "[| Notes A  {|Agent B,  Nonce PMS|} \<in> set evs;
+         Notes A' {|Agent B', Nonce PMS|} \<in> set evs;
+         evs \<in> tls |]
+      ==> A=A' & B=B'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule tls.induct, force, simp_all)
+(*ClientKeyExch*)
+apply (blast dest!: Notes_Crypt_parts_spies)
+done
+
+
+(**** Secrecy Theorems ****)
+
+(*Key compromise lemma needed to prove analz_image_keys.
+  No collection of keys can help the spy get new private keys.*)
+lemma analz_image_priK [rule_format]:
+     "evs \<in> tls
+      ==> \<forall>KK. (Key(priK B) \<in> analz (Key`KK Un (spies evs))) =
+          (priK B \<in> KK | B \<in> bad)"
+apply (erule tls.induct)
+apply (simp_all (no_asm_simp)
+		del: image_insert
+                add: image_Un [THEN sym]
+                     insert_Key_image Un_assoc [THEN sym])
+(*Fake*)
+apply spy_analz
+done
+
+
+(*slightly speeds up the big simplification below*)
+lemma range_sessionkeys_not_priK:
+     "KK <= range sessionK ==> priK B \<notin> KK"
+by blast
+
+
+(*Lemma for the trivial direction of the if-and-only-if*)
+lemma analz_image_keys_lemma:
+     "(X \<in> analz (G Un H)) --> (X \<in> analz H)  ==>
+      (X \<in> analz (G Un H))  =  (X \<in> analz H)"
+by (blast intro: analz_mono [THEN subsetD])
+
+(** Strangely, the following version doesn't work:
+\<forall>Z. (Nonce N \<in> analz (Key`(sessionK`Z) Un (spies evs))) =
+    (Nonce N \<in> analz (spies evs))"
+**)
+
+lemma analz_image_keys [rule_format]:
+     "evs \<in> tls ==>
+      \<forall>KK. KK <= range sessionK -->
+	      (Nonce N \<in> analz (Key`KK Un (spies evs))) =
+	      (Nonce N \<in> analz (spies evs))"
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (safe del: iffI)
+apply (safe del: impI iffI intro!: analz_image_keys_lemma)
+apply (simp_all (no_asm_simp)               (*faster*)
+                del: image_insert imp_disjL (*reduces blow-up*)
+		add: image_Un [THEN sym]  Un_assoc [THEN sym]
+		     insert_Key_singleton
+		     range_sessionkeys_not_priK analz_image_priK)
+apply (simp_all add: insert_absorb)
+(*Fake*)
+apply spy_analz
+done
+
+(*Knowing some session keys is no help in getting new nonces*)
+lemma analz_insert_key [simp]:
+     "evs \<in> tls ==>
+      Nonce N \<in> analz (insert (Key (sessionK z)) (spies evs)) =
+      (Nonce N \<in> analz (spies evs))"
+by (simp del: image_insert
+         add: insert_Key_singleton analz_image_keys)
+
+
+(*** Protocol goal: serverK(Na,Nb,M) and clientK(Na,Nb,M) remain secure ***)
+
+(** Some lemmas about session keys, comprising clientK and serverK **)
+
+
+(*Lemma: session keys are never used if PMS is fresh.
+  Nonces don't have to agree, allowing session resumption.
+  Converse doesn't hold; revealing PMS doesn't force the keys to be sent.
+  THEY ARE NOT SUITABLE AS SAFE ELIM RULES.*)
+lemma PMS_lemma:
+     "[| Nonce PMS \<notin> parts (spies evs);
+         K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role);
+         evs \<in> tls |]
+   ==> Key K \<notin> parts (spies evs) & (\<forall>Y. Crypt K Y \<notin> parts (spies evs))"
+apply (erule rev_mp, erule ssubst)
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp))
+(*Fake*)
+apply (blast intro: parts_insertI)
+(*SpyKeys*)
+apply blast
+(*Many others*)
+apply (force dest!: Notes_Crypt_parts_spies Notes_master_imp_Crypt_PMS)+
+done
+
+lemma PMS_sessionK_not_spied:
+     "[| Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) \<in> parts (spies evs);
+         evs \<in> tls |]
+      ==> Nonce PMS \<in> parts (spies evs)"
+by (blast dest: PMS_lemma)
+
+lemma PMS_Crypt_sessionK_not_spied:
+     "[| Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) Y
+           \<in> parts (spies evs);  evs \<in> tls |]
+      ==> Nonce PMS \<in> parts (spies evs)"
+by (blast dest: PMS_lemma)
+
+(*Write keys are never sent if M (MASTER SECRET) is secure.
+  Converse fails; betraying M doesn't force the keys to be sent!
+  The strong Oops condition can be weakened later by unicity reasoning,
+  with some effort.
+  NO LONGER USED: see clientK_not_spied and serverK_not_spied*)
+lemma sessionK_not_spied:
+     "[| \<forall>A. Says A Spy (Key (sessionK((NA,NB,M),role))) \<notin> set evs;
+         Nonce M \<notin> analz (spies evs);  evs \<in> tls |]
+      ==> Key (sessionK((NA,NB,M),role)) \<notin> parts (spies evs)"
+apply (erule rev_mp, erule rev_mp)
+apply (erule tls.induct, analz_mono_contra)
+apply (force, simp_all (no_asm_simp))
+(*Fake, SpyKeys*)
+apply blast+
+done
+
+
+(*If A sends ClientKeyExch to an honest B, then the PMS will stay secret.*)
+lemma Spy_not_see_PMS:
+     "[| Notes A {|Agent B, Nonce PMS|} \<in> set evs;
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
+      ==> Nonce PMS \<notin> analz (spies evs)"
+apply (erule rev_mp, erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp))
+(*Fake*)
+apply spy_analz
+(*SpyKeys*)
+apply force
+apply (simp_all add: insert_absorb) 
+(*ClientHello, ServerHello, ClientKeyExch: mostly freshness reasoning*)
+apply (blast dest: Notes_Crypt_parts_spies)
+apply (blast dest: Notes_Crypt_parts_spies)
+apply (blast dest: Notes_Crypt_parts_spies)
+(*ClientAccepts and ServerAccepts: because PMS \<notin> range PRF*)
+apply force+
+done
+
+
+(*If A sends ClientKeyExch to an honest B, then the MASTER SECRET
+  will stay secret.*)
+lemma Spy_not_see_MS:
+     "[| Notes A {|Agent B, Nonce PMS|} \<in> set evs;
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
+      ==> Nonce (PRF(PMS,NA,NB)) \<notin> analz (spies evs)"
+apply (erule rev_mp, erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp))
+(*Fake*)
+apply spy_analz
+(*SpyKeys: by secrecy of the PMS, Spy cannot make the MS*)
+apply (blast dest!: Spy_not_see_PMS)
+apply (simp_all add: insert_absorb)
+(*ClientAccepts and ServerAccepts: because PMS was already visible;
+  others, freshness etc.*)
+apply (blast dest: Notes_Crypt_parts_spies Spy_not_see_PMS 
+                   Notes_imp_knows_Spy [THEN analz.Inj])+
+done
+
+
+
+(*** Weakening the Oops conditions for leakage of clientK ***)
+
+(*If A created PMS then nobody else (except the Spy in replays)
+  would send a message using a clientK generated from that PMS.*)
+lemma Says_clientK_unique:
+     "[| Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) \<in> set evs;
+         Notes A {|Agent B, Nonce PMS|} \<in> set evs;
+         evs \<in> tls;  A' \<noteq> Spy |]
+      ==> A = A'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all)
+(*ClientKeyExch*)
+apply (blast dest!: PMS_Crypt_sessionK_not_spied)
+(*ClientFinished, ClientResume: by unicity of PMS*)
+apply (blast dest!: Notes_master_imp_Notes_PMS 
+             intro: Notes_unique_PMS [THEN conjunct1])+
+done
+
+
+(*If A created PMS and has not leaked her clientK to the Spy,
+  then it is completely secure: not even in parts!*)
+lemma clientK_not_spied:
+     "[| Notes A {|Agent B, Nonce PMS|} \<in> set evs;
+         Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) \<notin> set evs;
+         A \<notin> bad;  B \<notin> bad;
+         evs \<in> tls |]
+      ==> Key (clientK(Na,Nb,PRF(PMS,NA,NB))) \<notin> parts (spies evs)"
+apply (erule rev_mp, erule rev_mp)
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp))
+(*ClientKeyExch*)
+apply blast 
+(*SpyKeys*)
+apply (blast dest!: Spy_not_see_MS)
+(*ClientKeyExch*)
+apply (blast dest!: PMS_sessionK_not_spied)
+(*Oops*)
+apply (blast intro: Says_clientK_unique)
+done
+
+
+(*** Weakening the Oops conditions for leakage of serverK ***)
+
+(*If A created PMS for B, then nobody other than B or the Spy would
+  send a message using a serverK generated from that PMS.*)
+lemma Says_serverK_unique:
+     "[| Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) \<in> set evs;
+         Notes A {|Agent B, Nonce PMS|} \<in> set evs;
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad;  B' \<noteq> Spy |]
+      ==> B = B'"
+apply (erule rev_mp, erule rev_mp)
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all)
+(*ClientKeyExch*)
+apply (blast dest!: PMS_Crypt_sessionK_not_spied)
+(*ServerResume, ServerFinished: by unicity of PMS*)
+apply (blast dest!: Notes_master_imp_Crypt_PMS 
+             dest: Spy_not_see_PMS Notes_Crypt_parts_spies Crypt_unique_PMS)+
+done
+
+
+(*If A created PMS for B, and B has not leaked his serverK to the Spy,
+  then it is completely secure: not even in parts!*)
+lemma serverK_not_spied:
+     "[| Notes A {|Agent B, Nonce PMS|} \<in> set evs;
+         Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) \<notin> set evs;
+         A \<notin> bad;  B \<notin> bad;  evs \<in> tls |]
+      ==> Key (serverK(Na,Nb,PRF(PMS,NA,NB))) \<notin> parts (spies evs)"
+apply (erule rev_mp, erule rev_mp)
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp))
+(*Fake*)
+apply blast 
+(*SpyKeys*)
+apply (blast dest!: Spy_not_see_MS)
+(*ClientKeyExch*)
+apply (blast dest!: PMS_sessionK_not_spied)
+(*Oops*)
+apply (blast intro: Says_serverK_unique)
+done
+
+
+(*** Protocol goals: if A receives ServerFinished, then B is present
+     and has used the quoted values PA, PB, etc.  Note that it is up to A
+     to compare PA with what she originally sent.
+***)
+
+(*The mention of her name (A) in X assures A that B knows who she is.*)
+lemma TrustServerFinished [rule_format]:
+     "[| X = Crypt (serverK(Na,Nb,M))
+               (Hash{|Number SID, Nonce M,
+                      Nonce Na, Number PA, Agent A,
+                      Nonce Nb, Number PB, Agent B|});
+         M = PRF(PMS,NA,NB);
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
+      ==> Says B Spy (Key(serverK(Na,Nb,M))) \<notin> set evs -->
+          Notes A {|Agent B, Nonce PMS|} \<in> set evs -->
+          X \<in> parts (spies evs) --> Says B A X \<in> set evs"
+apply (erule ssubst)+
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp))
+(*Fake: the Spy doesn't have the critical session key!*)
+apply (blast dest: serverK_not_spied)
+(*ClientKeyExch*)
+apply (blast dest!: PMS_Crypt_sessionK_not_spied)
+done
+
+(*This version refers not to ServerFinished but to any message from B.
+  We don't assume B has received CertVerify, and an intruder could
+  have changed A's identity in all other messages, so we can't be sure
+  that B sends his message to A.  If CLIENT KEY EXCHANGE were augmented
+  to bind A's identity with PMS, then we could replace A' by A below.*)
+lemma TrustServerMsg [rule_format]:
+     "[| M = PRF(PMS,NA,NB);  evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
+      ==> Says B Spy (Key(serverK(Na,Nb,M))) \<notin> set evs -->
+          Notes A {|Agent B, Nonce PMS|} \<in> set evs -->
+          Crypt (serverK(Na,Nb,M)) Y \<in> parts (spies evs)  -->
+          (\<exists>A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) \<in> set evs)"
+apply (erule ssubst)
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp) add: ex_disj_distrib)
+(*Fake: the Spy doesn't have the critical session key!*)
+apply (blast dest: serverK_not_spied)
+(*ClientKeyExch*)
+apply (clarify, blast dest!: PMS_Crypt_sessionK_not_spied)
+(*ServerResume, ServerFinished: by unicity of PMS*)
+apply (blast dest!: Notes_master_imp_Crypt_PMS 
+             dest: Spy_not_see_PMS Notes_Crypt_parts_spies Crypt_unique_PMS)+
+done
+
+
+(*** Protocol goal: if B receives any message encrypted with clientK
+     then A has sent it, ASSUMING that A chose PMS.  Authentication is
+     assumed here; B cannot verify it.  But if the message is
+     ClientFinished, then B can then check the quoted values PA, PB, etc.
+***)
+
+lemma TrustClientMsg [rule_format]:
+     "[| M = PRF(PMS,NA,NB);  evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
+      ==> Says A Spy (Key(clientK(Na,Nb,M))) \<notin> set evs -->
+          Notes A {|Agent B, Nonce PMS|} \<in> set evs -->
+          Crypt (clientK(Na,Nb,M)) Y \<in> parts (spies evs) -->
+          Says A B (Crypt (clientK(Na,Nb,M)) Y) \<in> set evs"
+apply (erule ssubst)
+apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
+apply (force, simp_all (no_asm_simp))
+(*Fake: the Spy doesn't have the critical session key!*)
+apply (blast dest: clientK_not_spied)
+(*ClientKeyExch*)
+apply (blast dest!: PMS_Crypt_sessionK_not_spied)
+(*ClientFinished, ClientResume: by unicity of PMS*)
+apply (blast dest!: Notes_master_imp_Notes_PMS dest: Notes_unique_PMS)+
+done
+
+
+(*** Protocol goal: if B receives ClientFinished, and if B is able to
+     check a CertVerify from A, then A has used the quoted
+     values PA, PB, etc.  Even this one requires A to be uncompromised.
+ ***)
+lemma AuthClientFinished:
+     "[| M = PRF(PMS,NA,NB);
+         Says A Spy (Key(clientK(Na,Nb,M))) \<notin> set evs;
+         Says A' B (Crypt (clientK(Na,Nb,M)) Y) \<in> set evs;
+         certificate A KA \<in> parts (spies evs);
+         Says A'' B (Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|}))
+           \<in> set evs;
+         evs \<in> tls;  A \<notin> bad;  B \<notin> bad |]
+      ==> Says A B (Crypt (clientK(Na,Nb,M)) Y) \<in> set evs"
+by (blast intro!: TrustClientMsg UseCertVerify)
+
+(*22/9/97: loads in 622s, which is 10 minutes 22 seconds*)
+(*24/9/97: loads in 672s, which is 11 minutes 12 seconds [stronger theorems]*)
+(*29/9/97: loads in 481s, after removing Certificate from ClientKeyExch*)
+(*30/9/97: loads in 476s, after removing unused theorems*)
+(*30/9/97: loads in 448s, after fixing ServerResume*)
+
+(*08/9/97: loads in 189s (pike), after much reorganization,
+           back to 621s on albatross?*)
+
+(*10/2/99: loads in 139s (pike)
+           down to 433s on albatross*)
+
+(*5/5/01: conversion to Isar script
+	  loads in 137s (perch)
+          the last ML version loaded in 122s on perch, a 600MHz machine:
+		twice as fast as pike.  No idea why it's so much slower!
+	  The Isar script is slower still, perhaps because simp_all simplifies
+	  the assumptions be default.
+*)
 
 end
--- a/src/HOL/IsaMakefile	Mon May 07 19:19:41 2001 +0200
+++ b/src/HOL/IsaMakefile	Tue May 08 15:56:57 2001 +0200
@@ -307,7 +307,7 @@
   Auth/NS_Shared.thy Auth/OtwayRees.thy Auth/OtwayRees_AN.thy \
   Auth/OtwayRees_Bad.thy Auth/Public_lemmas.ML Auth/Public.thy Auth/ROOT.ML \
   Auth/Recur.thy Auth/Shared_lemmas.ML Auth/Shared.thy \
-  Auth/TLS.ML Auth/TLS.thy Auth/WooLam.thy \
+  Auth/TLS.thy Auth/WooLam.thy \
   Auth/Kerberos_BAN.ML Auth/Kerberos_BAN.thy \
   Auth/KerberosIV.ML Auth/KerberosIV.thy \
   Auth/Yahalom.thy Auth/Yahalom2.thy Auth/Yahalom_Bad.thy