Some X-symbols for <notin>, <noteq>, <forall>, <exists>
authorpaulson
Tue, 27 Feb 2001 16:13:23 +0100
changeset 11185 1b737b4c2108
parent 11184 10a307328d2c
child 11186 63f3e98df2a4
Some X-symbols for <notin>, <noteq>, <forall>, <exists> Streamlining of Yahalom proofs Removal of redundant proofs
src/HOL/Auth/KerberosIV.ML
src/HOL/Auth/KerberosIV.thy
src/HOL/Auth/Kerberos_BAN.ML
src/HOL/Auth/Kerberos_BAN.thy
src/HOL/Auth/OtwayRees.ML
src/HOL/Auth/OtwayRees.thy
src/HOL/Auth/OtwayRees_AN.ML
src/HOL/Auth/OtwayRees_AN.thy
src/HOL/Auth/OtwayRees_Bad.ML
src/HOL/Auth/OtwayRees_Bad.thy
src/HOL/Auth/Recur.ML
src/HOL/Auth/Recur.thy
src/HOL/Auth/Shared_lemmas.ML
src/HOL/Auth/TLS.ML
src/HOL/Auth/TLS.thy
src/HOL/Auth/WooLam.ML
src/HOL/Auth/WooLam.thy
src/HOL/Auth/Yahalom.ML
src/HOL/Auth/Yahalom.thy
src/HOL/Auth/Yahalom2.ML
src/HOL/Auth/Yahalom2.thy
src/HOL/Auth/Yahalom_Bad.ML
src/HOL/Auth/Yahalom_Bad.thy
--- a/src/HOL/Auth/KerberosIV.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/KerberosIV.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -63,8 +63,8 @@
 qed "AuthKeys_empty";
 
 Goalw [AuthKeys_def] 
- "(ALL A Tk akey Peer.              \
-\  ev ~= Says Kas A (Crypt (shrK A) {|akey, Agent Peer, Tk,      \
+ "(\\<forall>A Tk akey Peer.              \
+\  ev \\<noteq> Says Kas A (Crypt (shrK A) {|akey, Agent Peer, Tk,      \
 \             (Crypt (shrK Peer) {|Agent A, Agent Peer, akey, Tk|})|}))\ 
 \      ==> AuthKeys (ev # evs) = AuthKeys evs";
 by Auto_tac;
@@ -79,21 +79,21 @@
 qed "AuthKeys_insert";
 
 Goalw [AuthKeys_def] 
-   "K : AuthKeys \
+   "K \\<in> AuthKeys \
 \   (Says Kas A (Crypt (shrK A) {|Key K', Agent Peer, Number Tk, \
 \    (Crypt (shrK Peer) {|Agent A, Agent Peer, Key K', Number Tk|})|}) # evs) \
-\       ==> K = K' | K : AuthKeys evs";
+\       ==> K = K' | K \\<in> AuthKeys evs";
 by Auto_tac;
 qed "AuthKeys_simp";
 
 Goalw [AuthKeys_def]  
    "Says Kas A (Crypt (shrK A) {|Key K, Agent Tgs, Number Tk, \
-\    (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key K, Number Tk|})|}) : set evs \
-\       ==> K : AuthKeys evs";
+\    (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key K, Number Tk|})|}) \\<in> set evs \
+\       ==> K \\<in> AuthKeys evs";
 by Auto_tac;
 qed "AuthKeysI";
 
-Goalw [AuthKeys_def] "K : AuthKeys evs ==> Key K : used evs";
+Goalw [AuthKeys_def] "K \\<in> AuthKeys evs ==> Key K \\<in> used evs";
 by (Simp_tac 1);
 by (blast_tac (claset() addSEs spies_partsEs) 1);
 qed "AuthKeys_used";
@@ -103,18 +103,18 @@
 
 (*--For reasoning about the encrypted portion of message K3--*)
 Goal "Says Kas' A (Crypt KeyA {|AuthKey, Peer, Tk, AuthTicket|}) \
-\              : set evs ==> AuthTicket : parts (spies evs)";
+\              \\<in> set evs ==> AuthTicket \\<in> parts (spies evs)";
 by (blast_tac (claset() addSEs spies_partsEs) 1);
 qed "K3_msg_in_parts_spies";
 
 Goal "Says Kas A (Crypt KeyA {|AuthKey, Peer, Tk, AuthTicket|}) \
-\              : set evs ==> AuthKey : parts (spies evs)";
+\              \\<in> set evs ==> AuthKey \\<in> parts (spies evs)";
 by (blast_tac (claset() addSEs spies_partsEs) 1);
 qed "Oops_parts_spies1";
                               
 Goal "[| Says Kas A (Crypt KeyA {|Key AuthKey, Peer, Tk, AuthTicket|}) \
-\          : set evs ;\
-\        evs : kerberos |] ==> AuthKey ~: range shrK";
+\          \\<in> set evs ;\
+\        evs \\<in> kerberos |] ==> AuthKey \\<notin> range shrK";
 by (etac rev_mp 1);
 by (etac kerberos.induct 1);
 by Auto_tac;
@@ -122,25 +122,25 @@
 
 (*--For reasoning about the encrypted portion of message K5--*)
 Goal "Says Tgs' A (Crypt AuthKey {|ServKey, Agent B, Tt, ServTicket|})\
- \             : set evs ==> ServTicket : parts (spies evs)";
+ \             \\<in> set evs ==> ServTicket \\<in> parts (spies evs)";
 by (blast_tac (claset() addSEs spies_partsEs) 1);
 qed "K5_msg_in_parts_spies";
 
 Goal "Says Tgs A (Crypt AuthKey {|ServKey, Agent B, Tt, ServTicket|})\
-\                  : set evs ==> ServKey : parts (spies evs)";
+\                  \\<in> set evs ==> ServKey \\<in> parts (spies evs)";
 by (blast_tac (claset() addSEs spies_partsEs) 1);
 qed "Oops_parts_spies2";
 
 Goal "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}) \
-\          : set evs ;\
-\        evs : kerberos |] ==> ServKey ~: range shrK";
+\          \\<in> set evs ;\
+\        evs \\<in> kerberos |] ==> ServKey \\<notin> range shrK";
 by (etac rev_mp 1);
 by (etac kerberos.induct 1);
 by Auto_tac;
 qed "Oops_range_spies2";
 
-Goal "Says S A (Crypt K {|SesKey, B, TimeStamp, Ticket|}) : set evs \
-\     ==> Ticket : parts (spies evs)";
+Goal "Says S A (Crypt K {|SesKey, B, TimeStamp, Ticket|}) \\<in> set evs \
+\     ==> Ticket \\<in> parts (spies evs)";
 by (blast_tac (claset() addSEs spies_partsEs) 1);
 qed "Says_ticket_in_parts_spies";
 (*Replaces both K3_msg_in_parts_spies and K5_msg_in_parts_spies*)
@@ -156,44 +156,41 @@
 
 
 (*Spy never sees another agent's shared key! (unless it's lost at start)*)
-Goal "evs : kerberos ==> (Key (shrK A) : parts (spies evs)) = (A : bad)";
+Goal "evs \\<in> kerberos ==> (Key (shrK A) \\<in> parts (spies evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 by (ALLGOALS Blast_tac);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
-Goal "evs : kerberos ==> (Key (shrK A) : analz (spies evs)) = (A : bad)";
+Goal "evs \\<in> kerberos ==> (Key (shrK A) \\<in> analz (spies evs)) = (A \\<in> bad)";
 by (auto_tac (claset() addDs [impOfSubs analz_subset_parts], simpset()));
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
 
-Goal "[| Key (shrK A) : parts (spies evs);  evs : kerberos |] ==> A:bad";
+Goal "[| Key (shrK A) \\<in> parts (spies evs);  evs \\<in> kerberos |] ==> A:bad";
 by (blast_tac (claset() addDs [Spy_see_shrK]) 1);
 qed "Spy_see_shrK_D";
 bind_thm ("Spy_analz_shrK_D", analz_subset_parts RS subsetD RS Spy_see_shrK_D);
 AddSDs [Spy_see_shrK_D, Spy_analz_shrK_D];
 
 (*Nobody can have used non-existent keys!*)
-Goal "evs : kerberos ==>      \
-\     Key K ~: used evs --> K ~: keysFor (parts (spies evs))";
+Goal "evs \\<in> kerberos ==>      \
+\     Key K \\<notin> used evs --> K \\<notin> keysFor (parts (spies evs))";
 by (parts_induct_tac 1);
 (*Fake*)
-by (best_tac
-      (claset() addSDs [impOfSubs (parts_insert_subset_Un RS keysFor_mono)]
-               addIs  [impOfSubs analz_subset_parts]
-               addDs  [impOfSubs (analz_subset_parts RS keysFor_mono)]
-               addss  (simpset())) 1);
+by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
 (*Others*)
 by (ALLGOALS (blast_tac (claset() addSEs spies_partsEs)));
 qed_spec_mp "new_keys_not_used";
+Addsimps [new_keys_not_used];
 
+(*Earlier, \\<forall>protocol proofs declared this theorem.  
+  But Yahalom and Kerberos IV are the only ones that need it!*)
 bind_thm ("new_keys_not_analzd",
           [analz_subset_parts RS keysFor_mono,
            new_keys_not_used] MRS contra_subsetD);
 
-Addsimps [new_keys_not_used, new_keys_not_analzd];
-
 
 (*********************** REGULARITY LEMMAS ***********************)
 (*       concerning the form of items passed in messages         *)
@@ -201,9 +198,9 @@
 
 (*Describes the form of AuthKey, AuthTicket, and K sent by Kas*)
 Goal "[| Says Kas A (Crypt K {|Key AuthKey, Agent Peer, Tk, AuthTicket|}) \
-\          : set evs;                 \
-\        evs : kerberos |]             \
-\     ==> AuthKey ~: range shrK & AuthKey : AuthKeys evs & \
+\          \\<in> set evs;                 \
+\        evs \\<in> kerberos |]             \
+\     ==> AuthKey \\<notin> range shrK & AuthKey \\<in> AuthKeys evs & \
 \ AuthTicket = (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|} ) &\
 \            K = shrK A  & Peer = Tgs";
 by (etac rev_mp 1);
@@ -221,20 +218,20 @@
   Generalised to any session keys (both AuthKey and ServKey).
 *)
 Goal "[| Crypt (shrK Tgs_B) {|Agent A, Agent Tgs_B, Key SesKey, Number T|}\
-\           : parts (spies evs); Tgs_B ~: bad;\
-\        evs : kerberos |]    \
-\     ==> SesKey ~: range shrK";
+\           \\<in> parts (spies evs); Tgs_B \\<notin> bad;\
+\        evs \\<in> kerberos |]    \
+\     ==> SesKey \\<notin> range shrK";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 qed "SesKey_is_session_key";
 
 Goal "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}  \
-\          : parts (spies evs);                              \
-\        evs : kerberos |]                          \
+\          \\<in> parts (spies evs);                              \
+\        evs \\<in> kerberos |]                          \
 \     ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, \
 \                Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}|})  \
-\           : set evs";
+\           \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 (*Fake*)
@@ -244,9 +241,9 @@
 qed "A_trusts_AuthTicket";
 
 Goal "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}\
-\          : parts (spies evs);\
-\        evs : kerberos |]    \
-\     ==> AuthKey : AuthKeys evs";
+\          \\<in> parts (spies evs);\
+\        evs \\<in> kerberos |]    \
+\     ==> AuthKey \\<in> AuthKeys evs";
 by (ftac A_trusts_AuthTicket 1);
 by (assume_tac 1);
 by (simp_tac (simpset() addsimps [AuthKeys_def]) 1);
@@ -255,11 +252,11 @@
 
 (*Describes the form of ServKey, ServTicket and AuthKey sent by Tgs*)
 Goal "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})\
-\          : set evs; \
-\        evs : kerberos |]    \
-\  ==> B ~= Tgs & ServKey ~: range shrK & ServKey ~: AuthKeys evs &\
+\          \\<in> set evs; \
+\        evs \\<in> kerberos |]    \
+\  ==> B \\<noteq> Tgs & ServKey \\<notin> range shrK & ServKey \\<notin> AuthKeys evs &\
 \      ServTicket = (Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|} ) & \
-\      AuthKey ~: range shrK & AuthKey : AuthKeys evs";
+\      AuthKey \\<notin> range shrK & AuthKey \\<in> AuthKeys evs";
 by (etac rev_mp 1);
 by (etac kerberos.induct 1);
 by (ALLGOALS
@@ -277,10 +274,10 @@
 
 (*If a certain encrypted message appears then it originated with Kas*)
 Goal "[| Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|}  \
-\          : parts (spies evs);                              \
-\        A ~: bad;  evs : kerberos |]                        \
+\          \\<in> parts (spies evs);                              \
+\        A \\<notin> bad;  evs \\<in> kerberos |]                        \
 \     ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|})  \
-\           : set evs";
+\           \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 (*Fake*)
@@ -294,12 +291,12 @@
 
 (*If a certain encrypted message appears then it originated with Tgs*)
 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}     \
-\          : parts (spies evs);                                     \
-\        Key AuthKey ~: analz (spies evs);           \
-\        AuthKey ~: range shrK;                      \
-\        evs : kerberos |]         \
-\==> EX A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})\
-\      : set evs";
+\          \\<in> parts (spies evs);                                     \
+\        Key AuthKey \\<notin> analz (spies evs);           \
+\        AuthKey \\<notin> range shrK;                      \
+\        evs \\<in> kerberos |]         \
+\==> \\<exists>A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})\
+\      \\<in> set evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -312,10 +309,10 @@
 qed "A_trusts_K4";
 
 Goal "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|} \
-\          : parts (spies evs);          \
-\        A ~: bad;                       \
-\        evs : kerberos |]                \
-\   ==> AuthKey ~: range shrK &               \
+\          \\<in> parts (spies evs);          \
+\        A \\<notin> bad;                       \
+\        evs \\<in> kerberos |]                \
+\   ==> AuthKey \\<notin> range shrK &               \
 \       AuthTicket = Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -325,11 +322,11 @@
 
 (* This form holds also over an AuthTicket, but is not needed below.     *)
 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|} \
-\             : parts (spies evs); \
-\           Key AuthKey ~: analz (spies evs);  \
-\           evs : kerberos |]                                       \
-\        ==> ServKey ~: range shrK &  \
-\   (EX A. ServTicket = Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})";
+\             \\<in> parts (spies evs); \
+\           Key AuthKey \\<notin> analz (spies evs);  \
+\           evs \\<in> kerberos |]                                       \
+\        ==> ServKey \\<notin> range shrK &  \
+\   (\\<exists>A. ServTicket = Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -337,13 +334,13 @@
 qed "ServTicket_form";
 
 Goal "[| Says Kas' A (Crypt (shrK A) \
-\             {|Key AuthKey, Agent Tgs, Tk, AuthTicket|} ) : set evs; \
-\        evs : kerberos |]    \
-\     ==> AuthKey ~: range shrK & \
+\             {|Key AuthKey, Agent Tgs, Tk, AuthTicket|} ) \\<in> set evs; \
+\        evs \\<in> kerberos |]    \
+\     ==> AuthKey \\<notin> range shrK & \
 \         AuthTicket = \
 \                 Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}\
-\         | AuthTicket : analz (spies evs)";
-by (case_tac "A : bad" 1);
+\         | AuthTicket \\<in> analz (spies evs)";
+by (case_tac "A \\<in> bad" 1);
 by (force_tac (claset() addSDs [Says_imp_spies RS analz.Inj], simpset()) 1);
 by (forward_tac [Says_imp_spies RS parts.Inj] 1);
 by (blast_tac (claset() addSDs [AuthTicket_form]) 1);
@@ -351,13 +348,13 @@
 (* Essentially the same as AuthTicket_form *)
 
 Goal "[| Says Tgs' A (Crypt AuthKey \
-\             {|Key ServKey, Agent B, Tt, ServTicket|} ) : set evs; \
-\        evs : kerberos |]    \
-\     ==> ServKey ~: range shrK & \
-\         (EX A. ServTicket = \
+\             {|Key ServKey, Agent B, Tt, ServTicket|} ) \\<in> set evs; \
+\        evs \\<in> kerberos |]    \
+\     ==> ServKey \\<notin> range shrK & \
+\         (\\<exists>A. ServTicket = \
 \                 Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})  \
-\          | ServTicket : analz (spies evs)";
-by (case_tac "Key AuthKey : analz (spies evs)" 1);
+\          | ServTicket \\<in> analz (spies evs)";
+by (case_tac "Key AuthKey \\<in> analz (spies evs)" 1);
 by (blast_tac (claset() addDs [Says_imp_spies RS analz.Inj]) 1); 
 by (forward_tac [Says_imp_spies RS parts.Inj] 1);
 by (blast_tac (claset() addSDs [ServTicket_form]) 1);
@@ -372,10 +369,10 @@
    also Tgs in the place of B.                                     *)
 
 Goal "[| Crypt (shrK B)  {|Agent A,  Agent B,  Key SesKey, T|}        \
-\          : parts (spies evs);            \
+\          \\<in> parts (spies evs);            \
 \        Crypt (shrK B') {|Agent A', Agent B', Key SesKey, T'|}     \
-\          : parts (spies evs);  Key SesKey ~: analz (spies evs);   \
-\        evs : kerberos |]  \
+\          \\<in> parts (spies evs);  Key SesKey \\<notin> analz (spies evs);   \
+\        evs \\<in> kerberos |]  \
 \     ==> A=A' & B=B' & T=T'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -390,10 +387,10 @@
   A ServKey is encrypted by one and only one AuthKey.
 *)
 Goal "[| Crypt K  {|Key SesKey,  Agent B, T, Ticket|}        \
-\          : parts (spies evs);            \
+\          \\<in> parts (spies evs);            \
 \        Crypt K' {|Key SesKey,  Agent B', T', Ticket'|}     \
-\          : parts (spies evs);  Key SesKey ~: analz (spies evs);            \
-\        evs : kerberos |]  \
+\          \\<in> parts (spies evs);  Key SesKey \\<notin> analz (spies evs);            \
+\        evs \\<in> kerberos |]  \
 \     ==> K=K' & B=B' & T=T' & Ticket=Ticket'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -414,20 +411,20 @@
 
   Therefore, a goal like
 
-   "evs : kerberos \
-  \  ==> Key Kc ~: analz (spies evs) -->   \
-  \        (EX K' B' T' Ticket'. ALL K B T Ticket.                          \
+   "evs \\<in> kerberos \
+  \  ==> Key Kc \\<notin> analz (spies evs) -->   \
+  \        (\\<exists>K' B' T' Ticket'. \\<forall>K B T Ticket.                          \
   \         Crypt Kc {|Key K, Agent B, T, Ticket|}    \
-  \          : parts (spies evs) --> K=K' & B=B' & T=T' & Ticket=Ticket')";
+  \          \\<in> parts (spies evs) --> K=K' & B=B' & T=T' & Ticket=Ticket')";
 
   would fail on the K2 and K4 cases.
 *)
 
 Goal "[| Says Kas A                                          \
-\             (Crypt Ka {|Key AuthKey, Agent Tgs, Tk, X|}) : set evs;     \ 
+\             (Crypt Ka {|Key AuthKey, Agent Tgs, Tk, X|}) \\<in> set evs;     \ 
 \        Says Kas A'                                         \
-\             (Crypt Ka' {|Key AuthKey, Agent Tgs, Tk', X'|}) : set evs;   \
-\        evs : kerberos |] ==> A=A' & Ka=Ka' & Tk=Tk' & X=X'";
+\             (Crypt Ka' {|Key AuthKey, Agent Tgs, Tk', X'|}) \\<in> set evs;   \
+\        evs \\<in> kerberos |] ==> A=A' & Ka=Ka' & Tk=Tk' & X=X'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -437,10 +434,10 @@
 
 (* ServKey uniquely identifies the message from Tgs *)
 Goal "[| Says Tgs A                                             \
-\             (Crypt K {|Key ServKey, Agent B, Tt, X|}) : set evs; \ 
+\             (Crypt K {|Key ServKey, Agent B, Tt, X|}) \\<in> set evs; \ 
 \        Says Tgs A'                                                 \
-\             (Crypt K' {|Key ServKey, Agent B', Tt', X'|}) : set evs; \
-\        evs : kerberos |] ==> A=A' & B=B' & K=K' & Tt=Tt' & X=X'";
+\             (Crypt K' {|Key ServKey, Agent B', Tt', X'|}) \\<in> set evs; \
+\        evs \\<in> kerberos |] ==> A=A' & B=B' & K=K' & Tt=Tt' & X=X'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -458,8 +455,8 @@
 
 Goalw [KeyCryptKey_def]
  "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |}) \
-\             : set evs;    \
-\           evs : kerberos |] ==> KeyCryptKey AuthKey ServKey evs";
+\             \\<in> set evs;    \
+\           evs \\<in> kerberos |] ==> KeyCryptKey AuthKey ServKey evs";
 by (ftac Says_Tgs_message_form 1);
 by (assume_tac 1);
 by (Blast_tac 1);
@@ -468,7 +465,7 @@
 Goalw [KeyCryptKey_def]
    "KeyCryptKey AuthKey ServKey (Says S A X # evs) =                       \
 \    (Tgs = S &                                                            \
-\     (EX B tt. X = Crypt AuthKey        \
+\     (\\<exists>B tt. X = Crypt AuthKey        \
 \               {|Key ServKey, Agent B, tt,  \
 \                 Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} |}) \
 \    | KeyCryptKey AuthKey ServKey evs)";
@@ -480,7 +477,7 @@
 (*A fresh AuthKey cannot be associated with any other
   (with respect to a given trace). *)
 Goalw [KeyCryptKey_def]
- "[| Key AuthKey ~: used evs; evs : kerberos |] \
+ "[| Key AuthKey \\<notin> used evs; evs \\<in> kerberos |] \
 \        ==> ~ KeyCryptKey AuthKey ServKey evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -491,13 +488,13 @@
 (*A fresh ServKey cannot be associated with any other
   (with respect to a given trace). *)
 Goalw [KeyCryptKey_def]
- "Key ServKey ~: used evs ==> ~ KeyCryptKey AuthKey ServKey evs";
+ "Key ServKey \\<notin> used evs ==> ~ KeyCryptKey AuthKey ServKey evs";
 by (blast_tac (claset() addSEs spies_partsEs) 1);
 qed "Serv_fresh_not_KeyCryptKey";
 
 Goalw [KeyCryptKey_def]
  "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, tk|}\
-\             : parts (spies evs);  evs : kerberos |] \
+\             \\<in> parts (spies evs);  evs \\<in> kerberos |] \
 \        ==> ~ KeyCryptKey K AuthKey evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -511,9 +508,9 @@
 (*A secure serverkey cannot have been used to encrypt others*)
 Goalw [KeyCryptKey_def]
  "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} \
-\             : parts (spies evs);                     \
-\           Key ServKey ~: analz (spies evs);             \
-\           B ~= Tgs;  evs : kerberos |] \
+\             \\<in> parts (spies evs);                     \
+\           Key ServKey \\<notin> analz (spies evs);             \
+\           B \\<noteq> Tgs;  evs \\<in> kerberos |] \
 \        ==> ~ KeyCryptKey ServKey K evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -536,7 +533,7 @@
 
 (*Long term keys are not issued as ServKeys*)
 Goalw [KeyCryptKey_def]
- "evs : kerberos ==> ~ KeyCryptKey K (shrK A) evs";
+ "evs \\<in> kerberos ==> ~ KeyCryptKey K (shrK A) evs";
 by (parts_induct_tac 1);
 qed "shrK_not_KeyCryptKey";
 
@@ -544,13 +541,13 @@
   other key AuthKey.*)
 Goalw [KeyCryptKey_def]
  "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |}) \
-\      : set evs;                                         \
-\    AuthKey' ~= AuthKey;  evs : kerberos |]                      \
+\      \\<in> set evs;                                         \
+\    AuthKey' \\<noteq> AuthKey;  evs \\<in> kerberos |]                      \
 \ ==> ~ KeyCryptKey AuthKey' ServKey evs";
 by (blast_tac (claset() addDs [unique_ServKeys]) 1);
 qed "Says_Tgs_KeyCryptKey";
 
-Goal "[| KeyCryptKey AuthKey ServKey evs;  evs : kerberos |] \
+Goal "[| KeyCryptKey AuthKey ServKey evs;  evs \\<in> kerberos |] \
 \     ==> ~ KeyCryptKey ServKey K evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -572,29 +569,29 @@
 
 (*We take some pains to express the property
   as a logical equivalence so that the simplifier can apply it.*)
-Goal "P --> (Key K : analz (Key`KK Un H)) --> (K:KK | Key K : analz H)  \
+Goal "P --> (Key K \\<in> analz (Key`KK Un H)) --> (K:KK | Key K \\<in> analz H)  \
 \     ==>       \
-\     P --> (Key K : analz (Key`KK Un H)) = (K:KK | Key K : analz H)";
+\     P --> (Key K \\<in> analz (Key`KK Un H)) = (K:KK | Key K \\<in> analz H)";
 by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1);
 qed "Key_analz_image_Key_lemma";
 
-Goal "[| KeyCryptKey K K' evs; evs : kerberos |] \
-\     ==> Key K' : analz (insert (Key K) (spies evs))";
+Goal "[| KeyCryptKey K K' evs; evs \\<in> kerberos |] \
+\     ==> Key K' \\<in> analz (insert (Key K) (spies evs))";
 by (full_simp_tac (simpset() addsimps [KeyCryptKey_def]) 1);
 by (Clarify_tac 1);
 by (dresolve_tac [Says_imp_spies RS analz.Inj RS analz_insertI] 1);
 by Auto_tac;
 qed "KeyCryptKey_analz_insert";
 
-Goal "[| K : AuthKeys evs Un range shrK;  evs : kerberos |]  \
-\     ==> ALL SK. ~ KeyCryptKey SK K evs";
+Goal "[| K \\<in> AuthKeys evs Un range shrK;  evs \\<in> kerberos |]  \
+\     ==> \\<forall>SK. ~ KeyCryptKey SK K evs";
 by (asm_full_simp_tac (simpset() addsimps [KeyCryptKey_def]) 1);
 by (blast_tac (claset() addDs [Says_Tgs_message_form]) 1);
 qed "AuthKeys_are_not_KeyCryptKey";
 
-Goal "[| K ~: AuthKeys evs; \
-\        K ~: range shrK; evs : kerberos |]  \
-\     ==> ALL SK. ~ KeyCryptKey K SK evs";
+Goal "[| K \\<notin> AuthKeys evs; \
+\        K \\<notin> range shrK; evs \\<in> kerberos |]  \
+\     ==> \\<forall>SK. ~ KeyCryptKey K SK evs";
 by (asm_full_simp_tac (simpset() addsimps [KeyCryptKey_def]) 1);
 by (blast_tac (claset() addDs [Says_Tgs_message_form]) 1);
 qed "not_AuthKeys_not_KeyCryptKey";
@@ -613,16 +610,16 @@
     REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE, exE]
 		  ORELSE' hyp_subst_tac)];
 
-Goal "[| KK <= -(range shrK); Key K : analz (spies evs); evs: kerberos |]   \
-\     ==> Key K : analz (Key ` KK Un spies evs)";
+Goal "[| KK <= -(range shrK); Key K \\<in> analz (spies evs); evs \\<in> kerberos |]   \
+\     ==> Key K \\<in> analz (Key ` KK Un spies evs)";
 by (blast_tac (claset() addDs [impOfSubs analz_mono]) 1);
 qed "analz_mono_KK";
 
 (*For the Oops2 case of the next theorem*)
-Goal "[| evs : kerberos;  \
+Goal "[| evs \\<in> kerberos;  \
 \        Says Tgs A (Crypt AuthKey \
 \                    {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
-\          : set evs |] \
+\          \\<in> set evs |] \
 \     ==> ~ KeyCryptKey ServKey SK evs";
 by (blast_tac (claset() addDs [KeyCryptKeyI, KeyCryptKey_not_KeyCryptKey]) 1);
 qed "Oops2_not_KeyCryptKey";
@@ -633,11 +630,11 @@
 (* exploited as simplification laws for analz, and also "limit the damage" *)
 (* in case of loss of a key to the spy. See ESORICS98.                     *)
 (* [simplified by LCP]                                                     *)
-Goal "evs : kerberos ==>                                         \
-\     (ALL SK KK. KK <= -(range shrK) -->                   \
-\     (ALL K: KK. ~ KeyCryptKey K SK evs)   -->           \
-\     (Key SK : analz (Key`KK Un (spies evs))) =        \
-\     (SK : KK | Key SK : analz (spies evs)))";
+Goal "evs \\<in> kerberos ==>                                         \
+\     (\\<forall>SK KK. KK <= -(range shrK) -->                   \
+\     (\\<forall>K \\<in> KK. ~ KeyCryptKey K SK evs)   -->           \
+\     (Key SK \\<in> analz (Key`KK Un (spies evs))) =        \
+\     (SK \\<in> KK | Key SK \\<in> analz (spies evs)))";
 by (etac kerberos.induct 1);
 by analz_sees_tac;
 by (REPEAT_FIRST (rtac allI));
@@ -660,7 +657,7 @@
 by (blast_tac (claset() addEs spies_partsEs 
                         addSDs [AuthKey_not_KeyCryptKey]) 1);
 (*K5*)
-by (case_tac "Key ServKey : analz (spies evs5)" 1);
+by (case_tac "Key ServKey \\<in> analz (spies evs5)" 1);
 (*If ServKey is compromised then the result follows directly...*)
 by (asm_simp_tac 
      (simpset() addsimps [analz_insert_eq, 
@@ -677,10 +674,10 @@
 
 (* First simplification law for analz: no session keys encrypt  *)
 (* authentication keys or shared keys.                          *)
-Goal "[| evs : kerberos;  K : (AuthKeys evs) Un range shrK;      \
-\        SesKey ~: range shrK |]                                 \
-\     ==> Key K : analz (insert (Key SesKey) (spies evs)) = \
-\         (K = SesKey | Key K : analz (spies evs))";
+Goal "[| evs \\<in> kerberos;  K \\<in> (AuthKeys evs) Un range shrK;      \
+\        SesKey \\<notin> range shrK |]                                 \
+\     ==> Key K \\<in> analz (insert (Key SesKey) (spies evs)) = \
+\         (K = SesKey | Key K \\<in> analz (spies evs))";
 by (ftac AuthKeys_are_not_KeyCryptKey 1 THEN assume_tac 1);
 by (asm_full_simp_tac (analz_image_freshK_ss addsimps [Key_analz_image_Key]) 1);
 qed "analz_insert_freshK1";
@@ -688,9 +685,9 @@
 
 (* Second simplification law for analz: no service keys encrypt *)
 (* any other keys.					        *)
-Goal "[| evs : kerberos;  ServKey ~: (AuthKeys evs); ServKey ~: range shrK|]\
-\     ==> Key K : analz (insert (Key ServKey) (spies evs)) = \
-\         (K = ServKey | Key K : analz (spies evs))";
+Goal "[| evs \\<in> kerberos;  ServKey \\<notin> (AuthKeys evs); ServKey \\<notin> range shrK|]\
+\     ==> Key K \\<in> analz (insert (Key ServKey) (spies evs)) = \
+\         (K = ServKey | Key K \\<in> analz (spies evs))";
 by (ftac not_AuthKeys_not_KeyCryptKey 1 
     THEN assume_tac 1
     THEN assume_tac 1);
@@ -703,10 +700,10 @@
 Goal  
  "[| Says Tgs A    \
 \           (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
-\             : set evs;          \ 
-\           AuthKey ~= AuthKey'; AuthKey' ~: range shrK; evs : kerberos |]    \
-\       ==> Key ServKey : analz (insert (Key AuthKey') (spies evs)) =  \
-\               (ServKey = AuthKey' | Key ServKey : analz (spies evs))";
+\             \\<in> set evs;          \ 
+\           AuthKey \\<noteq> AuthKey'; AuthKey' \\<notin> range shrK; evs \\<in> kerberos |]    \
+\       ==> Key ServKey \\<in> analz (insert (Key AuthKey') (spies evs)) =  \
+\               (ServKey = AuthKey' | Key ServKey \\<in> analz (spies evs))";
 by (dres_inst_tac [("AuthKey'","AuthKey'")] Says_Tgs_KeyCryptKey 1);
 by (Blast_tac 1);
 by (assume_tac 1);
@@ -717,9 +714,9 @@
 (*a weakness of the protocol*)
 Goal "[| Says Tgs A    \
 \             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
-\          : set evs;          \ 
-\        Key AuthKey : analz (spies evs); evs : kerberos |]    \
-\     ==> Key ServKey : analz (spies evs)";
+\          \\<in> set evs;          \ 
+\        Key AuthKey \\<in> analz (spies evs); evs \\<in> kerberos |]    \
+\     ==> Key ServKey \\<in> analz (spies evs)";
 by (force_tac (claset() addDs [Says_imp_spies RS analz.Inj RS 
 			       analz.Decrypt RS analz.Fst],
 	       simpset()) 1);
@@ -729,10 +726,10 @@
 (********************** Guarantees for Kas *****************************)
 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, \
 \                     Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|}\
-\          : parts (spies evs); \
-\        Key ServKey ~: analz (spies evs);                          \
-\        B ~= Tgs; evs : kerberos |]                            \
-\     ==> ServKey ~: AuthKeys evs";
+\          \\<in> parts (spies evs); \
+\        Key ServKey \\<notin> analz (spies evs);                          \
+\        B \\<noteq> Tgs; evs \\<in> kerberos |]                            \
+\     ==> ServKey \\<notin> AuthKeys evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (asm_full_simp_tac (simpset() addsimps [AuthKeys_def]) 1);
@@ -745,13 +742,13 @@
 
 (** If Spy sees the Authentication Key sent in msg K2, then 
     the Key has expired  **)
-Goal "[| A ~: bad;  evs : kerberos |]           \
+Goal "[| A \\<notin> bad;  evs \\<in> kerberos |]           \
 \     ==> Says Kas A                             \
 \              (Crypt (shrK A)                      \
 \                 {|Key AuthKey, Agent Tgs, Number Tk,     \
 \         Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
-\           : set evs -->                 \
-\         Key AuthKey : analz (spies evs) -->                       \
+\           \\<in> set evs -->                 \
+\         Key AuthKey \\<in> analz (spies evs) -->                       \
 \         ExpirAuth Tk evs";
 by (etac kerberos.induct 1);
 by analz_sees_tac;
@@ -782,42 +779,38 @@
 
 Goal "[| Says Kas A                                             \
 \             (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})  \
-\          : set evs;                                \
+\          \\<in> set evs;                                \
 \        ~ ExpirAuth Tk evs;                         \
-\        A ~: bad;  evs : kerberos |]            \
-\     ==> Key AuthKey ~: analz (spies evs)";
+\        A \\<notin> bad;  evs \\<in> kerberos |]            \
+\     ==> Key AuthKey \\<notin> analz (spies evs)";
 by (ftac Says_Kas_message_form 1 THEN assume_tac 1);
 by (blast_tac (claset() addSDs [lemma]) 1);
 qed "Confidentiality_Kas";
 
 
-
-
-
-
 (********************** Guarantees for Tgs *****************************)
 
 (** If Spy sees the Service Key sent in msg K4, then 
     the Key has expired  **)
-Goal "[| A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]           \
-\  ==> Key AuthKey ~: analz (spies evs) --> \
+Goal "[| A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]           \
+\  ==> Key AuthKey \\<notin> analz (spies evs) --> \
 \      Says Tgs A            \
 \        (Crypt AuthKey                      \
 \           {|Key ServKey, Agent B, Number Tt,     \
 \             Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})\
-\        : set evs -->                 \
-\      Key ServKey : analz (spies evs) -->                       \
+\        \\<in> set evs -->                 \
+\      Key ServKey \\<in> analz (spies evs) -->                       \
 \      ExpirServ Tt evs";
 by (etac kerberos.induct 1);
-(*The Oops1 case is unusual: must simplify Authkey ~: analz (spies (ev#evs))
-  rather than weakening it to Authkey ~: analz (spies evs), for we then
-  conclude AuthKey ~= AuthKeya.*)
+(*The Oops1 case is unusual: must simplify Authkey \\<notin> analz (spies (ev#evs))
+  rather than weakening it to Authkey \\<notin> analz (spies evs), for we then
+  conclude AuthKey \\<noteq> AuthKeya.*)
 by (Clarify_tac 9);
 by analz_sees_tac;
 by (rotate_tac ~1 11);
 by (ALLGOALS 
     (asm_full_simp_tac 
-     (simpset() addsimps [less_SucI, 
+     (simpset() addsimps [less_SucI, new_keys_not_analzd,
 			  Says_Kas_message_form, Says_Tgs_message_form,
 			  analz_insert_eq, not_parts_not_analz, 
 			  analz_insert_freshK1, analz_insert_freshK2] 
@@ -826,9 +819,9 @@
 by (spy_analz_tac 1);
 (*K2*)
 by (blast_tac (claset() addSEs spies_partsEs
-            addIs [parts_insertI, impOfSubs analz_subset_parts, less_SucI]) 1);
+            addIs [parts_insertI, less_SucI]) 1);
 (*K4*)
-by (case_tac "A ~= Aa" 1);
+by (case_tac "A \\<noteq> Aa" 1);
 by (blast_tac (claset() addSEs spies_partsEs
                         addIs [less_SucI]) 1);
 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj RS parts.Fst, 
@@ -847,7 +840,7 @@
 			       Says_Kas_message_form, Says_Tgs_message_form] 
                         addIs  [less_SucI]) 2);
 (** Level 16 **)
-by (thin_tac "Says Aa Tgs ?X : set ?evs" 1);
+by (thin_tac "Says Aa Tgs ?X \\<in> set ?evs" 1);
 by (forward_tac [Says_imp_spies RS parts.Inj RS ServKey_notin_AuthKeysD] 1);
 by (assume_tac 1 THEN Blast_tac 1 THEN assume_tac 1);
 by (rotate_tac ~1 1);
@@ -863,11 +856,11 @@
 Goal 
  "[| Says Tgs A      \
 \             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
-\             : set evs;              \
-\           Key AuthKey ~: analz (spies evs);        \
+\             \\<in> set evs;              \
+\           Key AuthKey \\<notin> analz (spies evs);        \
 \           ~ ExpirServ Tt evs;                         \
-\           A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
-\        ==> Key ServKey ~: analz (spies evs)";
+\           A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
+\        ==> Key ServKey \\<notin> analz (spies evs)";
 by (ftac Says_Tgs_message_form 1 THEN assume_tac 1);
 by (blast_tac (claset() addDs [lemma]) 1);
 qed "Confidentiality_Tgs1";
@@ -876,13 +869,13 @@
 Goal 
  "[| Says Kas A                                             \
 \              (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})  \
-\             : set evs;                                \
+\             \\<in> set evs;                                \
 \           Says Tgs A      \
 \             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}) \
-\             : set evs;              \
+\             \\<in> set evs;              \
 \           ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
-\           A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
-\        ==> Key ServKey ~: analz (spies evs)";
+\           A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
+\        ==> Key ServKey \\<notin> analz (spies evs)";
 by (blast_tac (claset() addSDs [Confidentiality_Kas,
                                 Confidentiality_Tgs1]) 1);
 qed "Confidentiality_Tgs2";
@@ -897,13 +890,13 @@
 
 Goal
  "[| Says Kas A \
-\      (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) : set evs;\
+\      (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \\<in> set evs;\
 \    Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}     \
-\      : parts (spies evs);                                       \
-\    Key AuthKey ~: analz (spies evs);            \
-\    evs : kerberos |]         \
+\      \\<in> parts (spies evs);                                       \
+\    Key AuthKey \\<notin> analz (spies evs);            \
+\    evs \\<in> kerberos |]         \
 \==> Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})\
-\      : set evs";
+\      \\<in> set evs";
 by (ftac Says_Kas_message_form 1 THEN assume_tac 1);
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -919,12 +912,12 @@
 
 
 Goal "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}  \
-\          : parts (spies evs);                              \
+\          \\<in> parts (spies evs);                              \
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}     \
-\          : parts (spies evs);                                       \
+\          \\<in> parts (spies evs);                                       \
 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
-\     ==> Key ServKey ~: analz (spies evs)";
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
+\     ==> Key ServKey \\<notin> analz (spies evs)";
 by (dtac A_trusts_AuthKey 1);
 by (assume_tac 1);
 by (assume_tac 1);
@@ -939,10 +932,10 @@
 
 Goal
 "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})\
-\            : set evs; evs : kerberos|]  \
-\  ==> EX Tk. Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
+\            \\<in> set evs; evs \\<in> kerberos|]  \
+\  ==> \\<exists>Tk. Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
-\            : set evs";
+\            \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by Auto_tac;
@@ -952,10 +945,10 @@
 
 Goal
 "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})\
-\     : set evs; evs : kerberos|]  \
-\  ==> EX Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
+\     \\<in> set evs; evs \\<in> kerberos|]  \
+\  ==> \\<exists>Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
-\            : set evs   \
+\            \\<in> set evs   \
 \         & ServLife + Tt <= AuthLife + Tk)";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -965,12 +958,12 @@
 qed "K4_imp_K2_refined";
 
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}  \
-\          : parts (spies evs);  B ~= Tgs;  B ~: bad;       \
-\        evs : kerberos |]                        \
-\==> EX AuthKey. \
+\          \\<in> parts (spies evs);  B \\<noteq> Tgs;  B \\<notin> bad;       \
+\        evs \\<in> kerberos |]                        \
+\==> \\<exists>AuthKey. \
 \      Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt,  \
 \                  Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|}) \
-\      : set evs";
+\      \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
@@ -978,34 +971,34 @@
 qed "B_trusts_ServKey";
 
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
-\          : parts (spies evs);  B ~= Tgs;  B ~: bad;       \
-\        evs : kerberos |]                        \
-\  ==> EX AuthKey Tk. Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
+\          \\<in> parts (spies evs);  B \\<noteq> Tgs;  B \\<notin> bad;       \
+\        evs \\<in> kerberos |]                        \
+\  ==> \\<exists>AuthKey Tk. Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
-\            : set evs";
+\            \\<in> set evs";
 by (blast_tac (claset() addSDs [B_trusts_ServKey, K4_imp_K2]) 1);
 qed "B_trusts_ServTicket_Kas";
 
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
-\          : parts (spies evs); B ~= Tgs; B ~: bad;       \
-\        evs : kerberos |]                        \
-\  ==> EX AuthKey Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
+\          \\<in> parts (spies evs); B \\<noteq> Tgs; B \\<notin> bad;       \
+\        evs \\<in> kerberos |]                        \
+\  ==> \\<exists>AuthKey Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,\
 \          Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
-\            : set evs            \
+\            \\<in> set evs            \
 \          & ServLife + Tt <= AuthLife + Tk)";
 by (blast_tac (claset() addSDs [B_trusts_ServKey,K4_imp_K2_refined]) 1);
 qed "B_trusts_ServTicket_Kas_refined";
 
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
-\          : parts (spies evs); B ~= Tgs; B ~: bad;        \
-\        evs : kerberos |]                        \
-\==> EX Tk AuthKey.        \
+\          \\<in> parts (spies evs); B \\<noteq> Tgs; B \\<notin> bad;        \
+\        evs \\<in> kerberos |]                        \
+\==> \\<exists>Tk AuthKey.        \
 \    Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, \
 \                  Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
-\      : set evs         \ 
+\      \\<in> set evs         \ 
 \    & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,  \
 \                  Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|}) \
-\      : set evs";
+\      \\<in> set evs";
 by (ftac B_trusts_ServKey 1);
 by (etac exE 4);
 by (ftac K4_imp_K2 4);
@@ -1014,15 +1007,15 @@
 qed "B_trusts_ServTicket";
 
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
-\          : parts (spies evs); B ~= Tgs; B ~: bad;        \
-\        evs : kerberos |]                        \
-\==> EX Tk AuthKey.        \
+\          \\<in> parts (spies evs); B \\<noteq> Tgs; B \\<notin> bad;        \
+\        evs \\<in> kerberos |]                        \
+\==> \\<exists>Tk AuthKey.        \
 \    (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, \
 \                  Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})\
-\      : set evs         \ 
+\      \\<in> set evs         \ 
 \    & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,  \
 \                  Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|}) \
-\      : set evs         \
+\      \\<in> set evs         \
 \    & ServLife + Tt <= AuthLife + Tk)";
 by (ftac B_trusts_ServKey 1);
 by (etac exE 4);
@@ -1039,14 +1032,14 @@
 
 
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|} \
-\          : parts (spies evs);                                        \
+\          \\<in> parts (spies evs);                                        \
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|} \
-\          : parts (spies evs);                                         \
+\          \\<in> parts (spies evs);                                         \
 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
-\          : parts (spies evs);                     \
+\          \\<in> parts (spies evs);                     \
 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;     \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
-\     ==> Key ServKey ~: analz (spies evs)";
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
+\     ==> Key ServKey \\<notin> analz (spies evs)";
 by (ftac A_trusts_AuthKey 1);
 by (ftac Confidentiality_Kas 3);
 by (ftac B_trusts_ServTicket 6);
@@ -1070,10 +1063,10 @@
 
 (*Most general form -- only for refined model! *)
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}  \
-\          : parts (spies evs);                      \
+\          \\<in> parts (spies evs);                      \
 \        ~ ExpirServ Tt evs;                         \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
-\     ==> Key ServKey ~: analz (spies evs)";
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
+\     ==> Key ServKey \\<notin> analz (spies evs)";
 by (blast_tac (claset() addDs [B_trusts_ServTicket_refined,
 			       NotExpirServ_NotExpirAuth_refined, 
                                Confidentiality_Tgs2]) 1);
@@ -1088,12 +1081,12 @@
 
 (*Authenticity of ServKey for A: "A_trusts_ServKey"*)
 Goal "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|} \
-\          : parts (spies evs);                                     \
+\          \\<in> parts (spies evs);                                     \
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}   \
-\          : parts (spies evs);                                        \
-\        ~ ExpirAuth Tk evs; A ~: bad; evs : kerberos |]         \
+\          \\<in> parts (spies evs);                                        \
+\        ~ ExpirAuth Tk evs; A \\<notin> bad; evs \\<in> kerberos |]         \
 \==>Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})\
-\      : set evs";
+\      \\<in> set evs";
 by (ftac A_trusts_AuthKey 1 THEN assume_tac 1 THEN assume_tac 1);
 by (blast_tac (claset() addDs [Confidentiality_Auth_A, A_trusts_K4_bis]) 1);
 qed "A_trusts_ServKey"; 
@@ -1111,12 +1104,12 @@
 
 (*B checks authenticity of A: theorems "A_Authenticity", 
                                        "A_authenticity_refined" *)
-Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);  \
+Goal "[| Crypt ServKey {|Agent A, Number Ta|} \\<in> parts (spies evs);  \
 \        Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, \
-\                                    ServTicket|}) : set evs;       \
-\        Key ServKey ~: analz (spies evs);                \
-\        A ~: bad; B ~: bad; evs : kerberos |]   \
-\==> Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} : set evs";
+\                                    ServTicket|}) \\<in> set evs;       \
+\        Key ServKey \\<notin> analz (spies evs);                \
+\        A \\<notin> bad; B \\<notin> bad; evs \\<in> kerberos |]   \
+\==> Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} \\<in> set evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -1138,17 +1131,17 @@
 qed "Says_Auth";
 
 (*The second assumption tells B what kind of key ServKey is.*)
-Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);     \
+Goal "[| Crypt ServKey {|Agent A, Number Ta|} \\<in> parts (spies evs);     \
 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
-\          : parts (spies evs);                                         \
+\          \\<in> parts (spies evs);                                         \
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}  \ 
-\          : parts (spies evs);                                          \
+\          \\<in> parts (spies evs);                                          \
 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}  \
-\          : parts (spies evs);                                            \
+\          \\<in> parts (spies evs);                                            \
 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;  \
-\        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
+\        B \\<noteq> Tgs; A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos |]         \
 \  ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},\
-\                 Crypt ServKey {|Agent A, Number Ta|} |} : set evs";
+\                 Crypt ServKey {|Agent A, Number Ta|} |} \\<in> set evs";
 by (ftac Confidentiality_B 1);
 by (ftac B_trusts_ServKey 9);
 by (etac exE 12);
@@ -1158,13 +1151,13 @@
 qed "A_Authenticity";
 
 (*Stronger form in the refined model*)
-Goal "[| Crypt ServKey {|Agent A, Number Ta2|} : parts (spies evs);     \
+Goal "[| Crypt ServKey {|Agent A, Number Ta2|} \\<in> parts (spies evs);     \
 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
-\          : parts (spies evs);                                         \
+\          \\<in> parts (spies evs);                                         \
 \        ~ ExpirServ Tt evs;                                        \
-\        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
+\        B \\<noteq> Tgs; A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos |]         \
 \  ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},\
-\                 Crypt ServKey {|Agent A, Number Ta2|} |} : set evs";
+\                 Crypt ServKey {|Agent A, Number Ta2|} |} \\<in> set evs";
 by (ftac Confidentiality_B_refined 1);
 by (ftac B_trusts_ServKey 6);
 by (etac exE 9);
@@ -1176,12 +1169,12 @@
 
 (*A checks authenticity of B: theorem "B_authenticity"*)
 
-Goal "[| Crypt ServKey (Number Ta) : parts (spies evs);  \
+Goal "[| Crypt ServKey (Number Ta) \\<in> parts (spies evs);  \
 \        Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, \
-\                                    ServTicket|}) : set evs;       \
-\        Key ServKey ~: analz (spies evs);                \
-\        A ~: bad; B ~: bad; evs : kerberos |]   \
-\     ==> Says B A (Crypt ServKey (Number Ta)) : set evs";
+\                                    ServTicket|}) \\<in> set evs;       \
+\        Key ServKey \\<notin> analz (spies evs);                \
+\        A \\<notin> bad; B \\<notin> bad; evs \\<in> kerberos |]   \
+\     ==> Says B A (Crypt ServKey (Number Ta)) \\<in> set evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -1199,11 +1192,11 @@
 qed "Says_K6";
 
 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|}   \
-\          : parts (spies evs);    \
-\        Key AuthKey ~: analz (spies evs); AuthKey ~: range shrK;  \
-\        evs : kerberos |]              \
-\ ==> EX A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|})\
-\             : set evs";
+\          \\<in> parts (spies evs);    \
+\        Key AuthKey \\<notin> analz (spies evs); AuthKey \\<notin> range shrK;  \
+\        evs \\<in> kerberos |]              \
+\ ==> \\<exists>A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|})\
+\             \\<in> set evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -1212,14 +1205,14 @@
 by (Blast_tac 1);
 qed "K4_trustworthy";
 
-Goal "[| Crypt ServKey (Number Ta) : parts (spies evs);           \
+Goal "[| Crypt ServKey (Number Ta) \\<in> parts (spies evs);           \
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|} \
-\          : parts (spies evs);                                        \ 
+\          \\<in> parts (spies evs);                                        \ 
 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
-\          : parts (spies evs);                                          \
+\          \\<in> parts (spies evs);                                          \
 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
-\     ==> Says B A (Crypt ServKey (Number Ta)) : set evs";
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
+\     ==> Says B A (Crypt ServKey (Number Ta)) \\<in> set evs";
 by (ftac A_trusts_AuthKey 1);
 by (ftac Says_Kas_message_form 3);
 by (ftac Confidentiality_Kas 4);
@@ -1237,9 +1230,9 @@
 (***3. Parties' knowledge of session keys. A knows a session key if she
        used it to build a cipher.***)
 
-Goal "[| Says B A (Crypt ServKey (Number Ta)) : set evs;           \
-\        Key ServKey ~: analz (spies evs);                          \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
+Goal "[| Says B A (Crypt ServKey (Number Ta)) \\<in> set evs;           \
+\        Key ServKey \\<notin> analz (spies evs);                          \
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
 by (simp_tac (simpset() addsimps [Issues_def]) 1);
 by (rtac exI 1);
@@ -1262,41 +1255,41 @@
                         addIs [Says_K6]
                         addEs spies_partsEs) 1);
 qed "B_Knows_B_Knows_ServKey_lemma";
-(*Key ServKey ~: analz (spies evs) could be relaxed by Confidentiality_B
+(*Key ServKey \\<notin> analz (spies evs) could be relaxed by Confidentiality_B
   but this is irrelevant because B knows what he knows!                  *)
 
-Goal "[| Says B A (Crypt ServKey (Number Ta)) : set evs;           \
+Goal "[| Says B A (Crypt ServKey (Number Ta)) \\<in> set evs;           \
 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}\
-\           : parts (spies evs);\
+\           \\<in> parts (spies evs);\
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}\
-\           : parts (spies evs);\
+\           \\<in> parts (spies evs);\
 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
-\          : parts (spies evs);     \
+\          \\<in> parts (spies evs);     \
 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;              \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
 by (blast_tac (claset() addSDs [Confidentiality_B,
 	                       B_Knows_B_Knows_ServKey_lemma]) 1);
 qed "B_Knows_B_Knows_ServKey";
 
-Goal "[| Says B A (Crypt ServKey (Number Ta)) : set evs;           \
+Goal "[| Says B A (Crypt ServKey (Number Ta)) \\<in> set evs;           \
 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}\
-\           : parts (spies evs);\
+\           \\<in> parts (spies evs);\
 \        ~ ExpirServ Tt evs;            \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
 by (blast_tac (claset() addSDs [Confidentiality_B_refined,
 	                       B_Knows_B_Knows_ServKey_lemma]) 1);
 qed "B_Knows_B_Knows_ServKey_refined";
 
 
-Goal "[| Crypt ServKey (Number Ta) : parts (spies evs);           \
+Goal "[| Crypt ServKey (Number Ta) \\<in> parts (spies evs);           \
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|} \
-\          : parts (spies evs);                                        \ 
+\          \\<in> parts (spies evs);                                        \ 
 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
-\          : parts (spies evs);                                          \
+\          \\<in> parts (spies evs);                                          \
 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;                         \
-\        A ~: bad;  B ~: bad; B ~= Tgs; evs : kerberos |]            \
+\        A \\<notin> bad;  B \\<notin> bad; B \\<noteq> Tgs; evs \\<in> kerberos |]            \
 \     ==> B Issues A with (Crypt ServKey (Number Ta)) on evs";
 by (blast_tac (claset() addSDs [B_Authenticity, Confidentiality_Serv_A,
                                 B_Knows_B_Knows_ServKey_lemma]) 1);
@@ -1304,11 +1297,11 @@
 
 Goal "[| Says A Tgs     \
 \            {|AuthTicket, Crypt AuthKey {|Agent A, Number Ta|}, Agent B|}\
-\          : set evs;      \
-\        A ~: bad;  evs : kerberos |]         \
-\     ==> EX Tk. Says Kas A (Crypt (shrK A) \
+\          \\<in> set evs;      \
+\        A \\<notin> bad;  evs \\<in> kerberos |]         \
+\     ==> \\<exists>Tk. Says Kas A (Crypt (shrK A) \
 \                     {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \
-\                  : set evs";
+\                  \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
@@ -1318,15 +1311,15 @@
 qed "K3_imp_K2";
 
 Goal "[| Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}   \
-\          : parts (spies evs);                    \
+\          \\<in> parts (spies evs);                    \
 \        Says Kas A (Crypt (shrK A) \
 \                    {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \
-\        : set evs;    \
-\        Key AuthKey ~: analz (spies evs);       \
-\        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
+\        \\<in> set evs;    \
+\        Key AuthKey \\<notin> analz (spies evs);       \
+\        B \\<noteq> Tgs; A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos |]         \
 \  ==> Says Tgs A (Crypt AuthKey        \ 
 \                    {|Key ServKey, Agent B, Number Tt, ServTicket|})  \
-\        : set evs";      
+\        \\<in> set evs";      
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -1338,9 +1331,9 @@
 qed "K4_trustworthy'";
 
 Goal "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} \
-\          : set evs;       \
-\        Key ServKey ~: analz (spies evs);       \
-\        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
+\          \\<in> set evs;       \
+\        Key ServKey \\<notin> analz (spies evs);       \
+\        B \\<noteq> Tgs; A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos |]         \
 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
 by (simp_tac (simpset() addsimps [Issues_def]) 1);
 by (rtac exI 1);
@@ -1360,7 +1353,7 @@
 by (asm_full_simp_tac (simpset() addsimps [takeWhile_tail]) 1);
 (*Level 15: case study necessary because the assumption doesn't state
   the form of ServTicket. The guarantee becomes stronger.*)
-by (case_tac "Key AuthKey : analz (spies evs5)" 1);
+by (case_tac "Key AuthKey \\<in> analz (spies evs5)" 1);
 by (force_tac (claset() addDs [Says_imp_spies RS analz.Inj RS 
 			       analz.Decrypt RS analz.Fst],
 	       simpset()) 1);
@@ -1373,38 +1366,38 @@
 qed "A_Knows_A_Knows_ServKey_lemma";
 
 Goal "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} \
-\          : set evs;       \
+\          \\<in> set evs;       \
 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}\
-\          : parts (spies evs);\
+\          \\<in> parts (spies evs);\
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}\
-\          : parts (spies evs);                                        \
+\          \\<in> parts (spies evs);                                        \
 \        ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs;\
-\        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
+\        B \\<noteq> Tgs; A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos |]         \
 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
 by (blast_tac (claset() addSDs [Confidentiality_Serv_A,
 	                       A_Knows_A_Knows_ServKey_lemma]) 1);
 qed "A_Knows_A_Knows_ServKey";
 
-Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);     \
+Goal "[| Crypt ServKey {|Agent A, Number Ta|} \\<in> parts (spies evs);     \
 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
-\          : parts (spies evs);                                         \
+\          \\<in> parts (spies evs);                                         \
 \        Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}  \ 
-\          : parts (spies evs);                                          \
+\          \\<in> parts (spies evs);                                          \
 \        Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}  \
-\          : parts (spies evs);                                            \
+\          \\<in> parts (spies evs);                                            \
 \        ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs;  \
-\        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
+\        B \\<noteq> Tgs; A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos |]         \
 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
 by (blast_tac (claset() addDs [A_Authenticity, Confidentiality_B,
 	                       A_Knows_A_Knows_ServKey_lemma]) 1);
 qed "B_Knows_A_Knows_ServKey";
 
 
-Goal "[| Crypt ServKey {|Agent A, Number Ta|} : parts (spies evs);     \
+Goal "[| Crypt ServKey {|Agent A, Number Ta|} \\<in> parts (spies evs);     \
 \        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}       \
-\          : parts (spies evs);                                         \
+\          \\<in> parts (spies evs);                                         \
 \        ~ ExpirServ Tt evs;                                        \
-\        B ~= Tgs; A ~: bad;  B ~: bad;  evs : kerberos |]         \
+\        B \\<noteq> Tgs; A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos |]         \
 \  ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs";
 by (blast_tac (claset() addDs [A_Authenticity_refined, 
                                Confidentiality_B_refined,
--- a/src/HOL/Auth/KerberosIV.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/KerberosIV.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -19,7 +19,7 @@
 
 rules
   (*Tgs is secure --- we already know that Kas is secure*)
-  Tgs_not_bad "Tgs ~: bad"
+  Tgs_not_bad "Tgs \\<notin> bad"
   
 (*The current time is just the length of the trace!*)
 syntax
@@ -37,17 +37,17 @@
 constdefs
  (* AuthKeys are those contained in an AuthTicket *)
     AuthKeys :: event list => key set
-    "AuthKeys evs == {AuthKey. EX A Peer Tk. Says Kas A
+    "AuthKeys evs == {AuthKey. \\<exists>A Peer Tk. Says Kas A
                         (Crypt (shrK A) {|Key AuthKey, Agent Peer, Tk, 
                    (Crypt (shrK Peer) {|Agent A, Agent Peer, Key AuthKey, Tk|})
-                  |}) : set evs}"
+                  |}) \\<in> set evs}"
                       
  (* A is the true creator of X if she has sent X and X never appeared on
     the trace before this event. Recall that traces grow from head. *)
   Issues :: [agent , agent, msg, event list] => bool ("_ Issues _ with _ on _")
    "A Issues B with X on evs == 
-      EX Y. Says A B Y : set evs & X : parts {Y} &
-      X ~: parts (spies (takeWhile (% z. z  ~= Says A B Y) (rev evs)))"
+      \\<exists>Y. Says A B Y \\<in> set evs & X \\<in> parts {Y} &
+      X \\<notin> parts (spies (takeWhile (% z. z  \\<noteq> Says A B Y) (rev evs)))"
 
 
 consts
@@ -88,11 +88,11 @@
 constdefs 
   KeyCryptKey :: [key, key, event list] => bool
   "KeyCryptKey AuthKey ServKey evs ==
-     EX A B tt. 
+     \\<exists>A B tt. 
        Says Tgs A (Crypt AuthKey
                      {|Key ServKey, Agent B, tt,
                        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} |})
-         : set evs"
+         \\<in> set evs"
 
 consts
 
@@ -100,16 +100,15 @@
 inductive "kerberos"
   intrs 
         
-    Nil  "[]: kerberos"
+    Nil  "[] \\<in> kerberos"
 
-    Fake "[| evs: kerberos;  B ~= Spy;  
-             X: synth (analz (spies evs)) |]
-          ==> Says Spy B X  # evs : kerberos"
+    Fake "[| evsf \\<in> kerberos;  X \\<in> synth (analz (spies evsf)) |]
+          ==> Says Spy B X  # evsf \\<in> kerberos"
 
 (* FROM the initiator *)
-    K1   "[| evs1: kerberos |]
+    K1   "[| evs1 \\<in> kerberos |]
           ==> Says A Kas {|Agent A, Agent Tgs, Number (CT evs1)|} # evs1 
-          : kerberos"
+          \\<in> kerberos"
 
 (* Adding the timestamp serves to A in K3 to check that
    she doesn't get a reply too late. This kind of timeouts are ordinary. 
@@ -118,12 +117,12 @@
 (*---------------------------------------------------------------------*)
 
 (*FROM Kas *)
-    K2  "[| evs2: kerberos; Key AuthKey ~: used evs2;
-            Says A' Kas {|Agent A, Agent Tgs, Number Ta|} : set evs2 |]
+    K2  "[| evs2 \\<in> kerberos; Key AuthKey \\<notin> used evs2;
+            Says A' Kas {|Agent A, Agent Tgs, Number Ta|} \\<in> set evs2 |]
           ==> Says Kas A
                 (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number (CT evs2), 
                       (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, 
-                          Number (CT evs2)|})|}) # evs2 : kerberos"
+                          Number (CT evs2)|})|}) # evs2 \\<in> kerberos"
 (* 
   The internal encryption builds the AuthTicket.
   The timestamp doesn't change inside the two encryptions: the external copy
@@ -134,15 +133,15 @@
 (*---------------------------------------------------------------------*)
 
 (* FROM the initiator *)
-    K3  "[| evs3: kerberos; 
-            Says A Kas {|Agent A, Agent Tgs, Number Ta|} : set evs3;
+    K3  "[| evs3 \\<in> kerberos; 
+            Says A Kas {|Agent A, Agent Tgs, Number Ta|} \\<in> set evs3;
             Says Kas' A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, 
-              AuthTicket|}) : set evs3; 
+              AuthTicket|}) \\<in> set evs3; 
             RecentResp Tk Ta
          |]
           ==> Says A Tgs {|AuthTicket, 
                            (Crypt AuthKey {|Agent A, Number (CT evs3)|}), 
-                           Agent B|} # evs3 : kerberos"
+                           Agent B|} # evs3 \\<in> kerberos"
 (*The two events amongst the premises allow A to accept only those AuthKeys 
   that are not issued late. *)
 
@@ -153,12 +152,12 @@
    specification. Adding it strengthens the guarantees assessed by the 
    protocol. Theorems that exploit it have the suffix `_refined'
 *) 
-    K4  "[| evs4: kerberos; Key ServKey ~: used evs4; B ~= Tgs; 
+    K4  "[| evs4 \\<in> kerberos; Key ServKey \\<notin> used evs4; B \\<noteq> Tgs; 
             Says A' Tgs {|
              (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey,
 				 Number Tk|}),
              (Crypt AuthKey {|Agent A, Number Ta1|}), Agent B|}
-	        : set evs4;
+	        \\<in> set evs4;
             ~ ExpirAuth Tk evs4;
             ~ ExpirAutc Ta1 evs4; 
             ServLife + (CT evs4) <= AuthLife + Tk
@@ -167,7 +166,7 @@
                 (Crypt AuthKey {|Key ServKey, Agent B, Number (CT evs4),  
 			       Crypt (shrK B) {|Agent A, Agent B, Key ServKey,
 		 			        Number (CT evs4)|} |})
-	        # evs4 : kerberos"
+	        # evs4 \\<in> kerberos"
 (* Tgs creates a new session key per each request for a service, without 
    checking if there is still a fresh one for that service.
    The cipher under Tgs' key is the AuthTicket, the cipher under B's key
@@ -179,56 +178,56 @@
 (*---------------------------------------------------------------------*)
 
 (* FROM the initiator *)
-    K5  "[| evs5: kerberos;  
+    K5  "[| evs5 \\<in> kerberos;  
             Says A Tgs 
                 {|AuthTicket, (Crypt AuthKey {|Agent A, Number Ta1|} ),
 		  Agent B|}
-              : set evs5;
+              \\<in> set evs5;
             Says Tgs' A 
              (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|} ) 
-                : set evs5;
+                \\<in> set evs5;
             RecentResp Tt Ta1 |]
           ==> Says A B {|ServTicket,
 			 Crypt ServKey {|Agent A, Number (CT evs5)|} |}
-               # evs5 : kerberos"
+               # evs5 \\<in> kerberos"
 (* Checks similar to those in K3. *)
 
 (*---------------------------------------------------------------------*)
 
 (* FROM the responder*)
-     K6  "[| evs6: kerberos;
+     K6  "[| evs6 \\<in> kerberos;
             Says A' B {|           
               (Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|} ),
               (Crypt ServKey {|Agent A, Number Ta2|} )|}
-            : set evs6;
+            \\<in> set evs6;
             ~ ExpirServ Tt evs6;
             ~ ExpirAutc Ta2 evs6
          |]
           ==> Says B A (Crypt ServKey (Number Ta2) )
-               # evs6 : kerberos"
+               # evs6 \\<in> kerberos"
 (* Checks similar to those in K4. *)
 
 (*---------------------------------------------------------------------*)
 
 (* Leaking an AuthKey... *)
-    Oops1 "[| evsO1: kerberos;  A ~= Spy;
+    Oops1 "[| evsO1 \\<in> kerberos;  A \\<noteq> Spy;
               Says Kas A
                 (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, 
-                                  AuthTicket|})  : set evsO1;
+                                  AuthTicket|})  \\<in> set evsO1;
               ExpirAuth Tk evsO1 |]
           ==> Says A Spy {|Agent A, Agent Tgs, Number Tk, Key AuthKey|} 
-               # evsO1 : kerberos"
+               # evsO1 \\<in> kerberos"
 
 (*---------------------------------------------------------------------*)
 
 (*Leaking a ServKey... *)
-    Oops2 "[| evsO2: kerberos;  A ~= Spy;
+    Oops2 "[| evsO2 \\<in> kerberos;  A \\<noteq> Spy;
               Says Tgs A 
                 (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
-                   : set evsO2;
+                   \\<in> set evsO2;
               ExpirServ Tt evsO2 |]
           ==> Says A Spy {|Agent A, Agent B, Number Tt, Key ServKey|} 
-               # evsO2 : kerberos"
+               # evsO2 \\<in> kerberos"
 
 (*---------------------------------------------------------------------*)
 
--- a/src/HOL/Auth/Kerberos_BAN.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Kerberos_BAN.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -23,9 +23,9 @@
 
 
 (*A "possibility property": there are traces that reach the end.*)
-Goal "EX Timestamp K. EX evs: kerberos_ban.    \
+Goal "\\<exists>Timestamp K. \\<exists>evs \\<in> kerberos_ban.    \
 \            Says B A (Crypt K (Number Timestamp)) \
-\                 : set evs";
+\                 \\<in> set evs";
 by (cut_facts_tac [SesKeyLife_LB] 1);
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (kerberos_ban.Nil RS kerberos_ban.Kb1 RS kerberos_ban.Kb2 RS 
@@ -39,17 +39,17 @@
 (**** Inductive proofs about kerberos_ban ****)
 
 (*Forwarding Lemma for reasoning about the encrypted portion of message Kb3*)
-Goal "Says S A (Crypt KA {|Timestamp, B, K, X|}) : set evs \
-\             ==> X : parts (spies evs)";
+Goal "Says S A (Crypt KA {|Timestamp, B, K, X|}) \\<in> set evs \
+\             ==> X \\<in> parts (spies evs)";
 by (Blast_tac 1);
 qed "Kb3_msg_in_parts_spies";
                               
-Goal "Says Server A (Crypt (shrK A) {|Timestamp, B, K, X|}) : set evs \
-\        ==> K : parts (spies evs)";
+Goal "Says Server A (Crypt (shrK A) {|Timestamp, B, K, X|}) \\<in> set evs \
+\        ==> K \\<in> parts (spies evs)";
 by (Blast_tac 1);
 qed "Oops_parts_spies";
 
-(*For proving the easier theorems about X ~: parts (spies evs).*)
+(*For proving the easier theorems about X \\<notin> parts (spies evs).*)
 fun parts_induct_tac i = 
     etac kerberos_ban.induct i  THEN 
     ftac Oops_parts_spies (i+6)  THEN
@@ -58,20 +58,20 @@
 
 
 (*Spy never sees another agent's shared key! (unless it's bad at start)*)
-Goal "evs : kerberos_ban ==> (Key (shrK A) : parts (spies evs)) = (A : bad)";
+Goal "evs \\<in> kerberos_ban ==> (Key (shrK A) \\<in> parts (spies evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
 
-Goal "evs : kerberos_ban ==> (Key (shrK A) : analz (spies evs)) = (A : bad)";
+Goal "evs \\<in> kerberos_ban ==> (Key (shrK A) \\<in> analz (spies evs)) = (A \\<in> bad)";
 by Auto_tac;
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
 
-Goal  "[| Key (shrK A) : parts (spies evs);       \
-\               evs : kerberos_ban |] ==> A:bad";
+Goal  "[| Key (shrK A) \\<in> parts (spies evs);       \
+\               evs \\<in> kerberos_ban |] ==> A:bad";
 by (blast_tac (claset() addDs [Spy_see_shrK]) 1);
 qed "Spy_see_shrK_D";
 
@@ -80,28 +80,22 @@
 
 
 (*Nobody can have used non-existent keys!*)
-Goal "evs : kerberos_ban ==>      \
-\      Key K ~: used evs --> K ~: keysFor (parts (spies evs))";
+Goal "evs \\<in> kerberos_ban ==>      \
+\      Key K \\<notin> used evs --> K \\<notin> keysFor (parts (spies evs))";
 by (parts_induct_tac 1);
 (*Fake*)
 by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
 (*Kb2, Kb3, Kb4*)
 by (ALLGOALS Blast_tac);
 qed_spec_mp "new_keys_not_used";
-
-bind_thm ("new_keys_not_analzd",
-          [analz_subset_parts RS keysFor_mono,
-           new_keys_not_used] MRS contra_subsetD);
-
-Addsimps [new_keys_not_used, new_keys_not_analzd];
-
+Addsimps [new_keys_not_used];
 
 (** Lemmas concerning the form of items passed in messages **)
 
 (*Describes the form of K, X and K' when the Server sends this message.*)
 Goal "[| Says Server A (Crypt K' {|Number Ts, Agent B, Key K, X|})  \
-\        : set evs; evs : kerberos_ban |]                           \
-\     ==> K ~: range shrK &                                         \
+\        \\<in> set evs; evs \\<in> kerberos_ban |]                           \
+\     ==> K \\<notin> range shrK &                                         \
 \         X = (Crypt (shrK B) {|Number Ts, Agent A, Key K|}) &      \
 \         K' = shrK A";
 by (etac rev_mp 1);
@@ -116,10 +110,10 @@
   This shows implicitly the FRESHNESS OF THE SESSION KEY to A
 *)
 Goal "[| Crypt (shrK A) {|Number Ts, Agent B, Key K, X|} \
-\          : parts (spies evs);                          \
-\        A ~: bad;  evs : kerberos_ban |]                \
+\          \\<in> parts (spies evs);                          \
+\        A \\<notin> bad;  evs \\<in> kerberos_ban |]                \
 \      ==> Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) \
-\            : set evs";
+\            \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Blast_tac 1);
@@ -128,12 +122,12 @@
 
 (*If the TICKET appears then it originated with the Server*)
 (*FRESHNESS OF THE SESSION KEY to B*)
-Goal "[| Crypt (shrK B) {|Number Ts, Agent A, Key K|} : parts (spies evs); \
-\        B ~: bad;  evs : kerberos_ban |]                        \
+Goal "[| Crypt (shrK B) {|Number Ts, Agent A, Key K|} \\<in> parts (spies evs); \
+\        B \\<notin> bad;  evs \\<in> kerberos_ban |]                        \
 \      ==> Says Server A                                         \
 \           (Crypt (shrK A) {|Number Ts, Agent B, Key K,                   \
 \                         Crypt (shrK B) {|Number Ts, Agent A, Key K|}|})  \
-\          : set evs";
+\          \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Blast_tac 1);
@@ -144,11 +138,11 @@
   OR     reduces it to the Fake case.
   Use Says_Server_message_form if applicable.*)
 Goal "[| Says S A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})     \
-\           : set evs;                                                  \
-\        evs : kerberos_ban |]                                          \
-\==> (K ~: range shrK & X = (Crypt (shrK B) {|Number Ts, Agent A, Key K|}))\
-\         | X : analz (spies evs)";
-by (case_tac "A : bad" 1);
+\           \\<in> set evs;                                                  \
+\        evs \\<in> kerberos_ban |]                                          \
+\==> (K \\<notin> range shrK & X = (Crypt (shrK B) {|Number Ts, Agent A, Key K|}))\
+\         | X \\<in> analz (spies evs)";
+by (case_tac "A \\<in> bad" 1);
 by (fast_tac (claset() addSDs [Says_imp_spies RS analz.Inj]
                       addss (simpset())) 1);
 by (forward_tac [Says_imp_spies RS parts.Inj] 1);
@@ -167,8 +161,8 @@
 (****
  The following is to prove theorems of the form
 
-  Key K : analz (insert (Key KAB) (spies evs)) ==>
-  Key K : analz (spies evs)
+  Key K \\<in> analz (insert (Key KAB) (spies evs)) ==>
+  Key K \\<in> analz (spies evs)
 
  A more general formula must be proved inductively.
 
@@ -177,10 +171,10 @@
 
 (** Session keys are not used to encrypt other session keys **)
 
-Goal "evs : kerberos_ban ==>                          \
-\  ALL K KK. KK <= - (range shrK) -->                 \
-\         (Key K : analz (Key`KK Un (spies evs))) =  \
-\         (K : KK | Key K : analz (spies evs))";
+Goal "evs \\<in> kerberos_ban ==>                          \
+\  \\<forall>K KK. KK <= - (range shrK) -->                 \
+\         (Key K \\<in> analz (Key`KK Un (spies evs))) =  \
+\         (K \\<in> KK | Key K \\<in> analz (spies evs))";
 by (etac kerberos_ban.induct 1);
 by analz_spies_tac;
 by (REPEAT_FIRST (resolve_tac [allI, impI]));
@@ -192,9 +186,9 @@
 qed_spec_mp "analz_image_freshK";
 
 
-Goal "[| evs : kerberos_ban;  KAB ~: range shrK |] ==>     \
-\     Key K : analz (insert (Key KAB) (spies evs)) =       \
-\     (K = KAB | Key K : analz (spies evs))";
+Goal "[| evs \\<in> kerberos_ban;  KAB \\<notin> range shrK |] ==>     \
+\     Key K \\<in> analz (insert (Key KAB) (spies evs)) =       \
+\     (K = KAB | Key K \\<in> analz (spies evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);
 qed "analz_insert_freshK";
 
@@ -202,10 +196,10 @@
 (** The session key K uniquely identifies the message **)
 
 Goal "[| Says Server A                                    \
-\          (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) : set evs; \ 
+\          (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) \\<in> set evs; \ 
 \        Says Server A'                                   \
-\         (Crypt (shrK A') {|Number Ts', Agent B', Key K, X'|}) : set evs;\
-\        evs : kerberos_ban |] ==> A=A' & Ts=Ts' & B=B' & X = X'";
+\         (Crypt (shrK A') {|Number Ts', Agent B', Key K, X'|}) \\<in> set evs;\
+\        evs \\<in> kerberos_ban |] ==> A=A' & Ts=Ts' & B=B' & X = X'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -218,12 +212,12 @@
     if the spy could see it!
 **)
 
-Goal "[| A ~: bad;  B ~: bad;  evs : kerberos_ban |]           \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos_ban |]           \
 \ ==> Says Server A                                            \
 \         (Crypt (shrK A) {|Number Ts, Agent B, Key K,         \
 \                           Crypt (shrK B) {|Number Ts, Agent A, Key K|}|})\
-\        : set evs -->                                         \
-\     Key K : analz (spies evs) --> Expired Ts evs"; 
+\        \\<in> set evs -->                                         \
+\     Key K \\<in> analz (spies evs) --> Expired Ts evs"; 
 by (etac kerberos_ban.induct 1);
 by analz_spies_tac;
 by (ALLGOALS
@@ -237,7 +231,7 @@
 by (spy_analz_tac 1);
 (**LEVEL 6 **)
 (*Kb3*)
-by (case_tac "Aa : bad" 1);
+by (case_tac "Aa \\<in> bad" 1);
 by (blast_tac (claset() addDs [A_trusts_K_by_Kb2, unique_session_keys]) 2);
 by (blast_tac (claset() addDs [Says_imp_spies RS analz.Inj,
                                Crypt_Spy_analz_bad, analz.Fst, analz.Snd]
@@ -250,25 +244,25 @@
                      as long as they have NOT EXPIRED
 **)
 Goal "[| Says Server A                                           \
-\         (Crypt K' {|Number T, Agent B, Key K, X|}) : set evs;  \
+\         (Crypt K' {|Number T, Agent B, Key K, X|}) \\<in> set evs;  \
 \        ~ Expired T evs;                                        \
-\        A ~: bad;  B ~: bad;  evs : kerberos_ban                \
-\     |] ==> Key K ~: analz (spies evs)";
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos_ban                \
+\     |] ==> Key K \\<notin> analz (spies evs)";
 by (ftac Says_Server_message_form 1 THEN assume_tac 1);
 by (blast_tac (claset() addIs [lemma2]) 1);
 qed "Confidentiality_S";
 
 (**** THE COUNTERPART OF CONFIDENTIALITY 
-      [|...; Expired Ts evs; ...|] ==> Key K : analz (spies evs)
+      [|...; Expired Ts evs; ...|] ==> Key K \\<in> analz (spies evs)
       WOULD HOLD ONLY IF AN OOPS OCCURRED! ---> Nothing to prove!   ****)
 
 
 (** CONFIDENTIALITY for ALICE: **)
 (** Also A_trusts_K_by_Kb2 RS Confidentiality_S **)
-Goal "[| Crypt (shrK A) {|Number T, Agent B, Key K, X|} : parts (spies evs);\
+Goal "[| Crypt (shrK A) {|Number T, Agent B, Key K, X|} \\<in> parts (spies evs);\
 \        ~ Expired T evs;          \
-\        A ~: bad;  B ~: bad;  evs : kerberos_ban                \
-\     |] ==> Key K ~: analz (spies evs)";
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos_ban                \
+\     |] ==> Key K \\<notin> analz (spies evs)";
 by (blast_tac (claset() addSDs [A_trusts_K_by_Kb2, Confidentiality_S]) 1);
 qed "Confidentiality_A";
 
@@ -276,21 +270,21 @@
 (** CONFIDENTIALITY for BOB: **)
 (** Also B_trusts_K_by_Kb3 RS Confidentiality_S **)
 Goal "[| Crypt (shrK B) {|Number Tk, Agent A, Key K|} \
-\         : parts (spies evs);              \
+\         \\<in> parts (spies evs);              \
 \       ~ Expired Tk evs;          \
-\       A ~: bad;  B ~: bad;  evs : kerberos_ban                \
-\     |] ==> Key K ~: analz (spies evs)";             
+\       A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos_ban                \
+\     |] ==> Key K \\<notin> analz (spies evs)";             
 by (blast_tac (claset() addSDs [B_trusts_K_by_Kb3, 
                                 Confidentiality_S]) 1);
 qed "Confidentiality_B";
 
 
-Goal "[| B ~: bad;  evs : kerberos_ban |]                        \
-\     ==> Key K ~: analz (spies evs) -->                    \
+Goal "[| B \\<notin> bad;  evs \\<in> kerberos_ban |]                        \
+\     ==> Key K \\<notin> analz (spies evs) -->                    \
 \         Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) \
-\         : set evs -->                                             \
-\         Crypt K (Number Ta) : parts (spies evs) -->        \
-\         Says B A (Crypt K (Number Ta)) : set evs";
+\         \\<in> set evs -->                                             \
+\         Crypt K (Number Ta) \\<in> parts (spies evs) -->        \
+\         Says B A (Crypt K (Number Ta)) \\<in> set evs";
 by (etac kerberos_ban.induct 1);
 by (ftac Says_S_message_form 5 THEN assume_tac 5);     
 by (dtac Kb3_msg_in_parts_spies 5);
@@ -302,12 +296,12 @@
 by (Clarify_tac 1);
 (*
 Subgoal 1: contradiction from the assumptions  
-Key K ~: used evs2  and Crypt K (Number Ta) : parts (spies evs2)
+Key K \\<notin> used evs2  and Crypt K (Number Ta) \\<in> parts (spies evs2)
 *)
 by (dtac Crypt_imp_invKey_keysFor 1);
 by (Asm_full_simp_tac 1);
 (* the two tactics above detect the contradiction*)
-by (case_tac "Ba : bad" 1);  (*splits up the subgoal by the stated case*)
+by (case_tac "Ba \\<in> bad" 1);  (*splits up the subgoal by the stated case*)
 by (blast_tac (claset() addDs [Says_imp_spies RS parts.Inj RS parts.Fst RS 
                               B_trusts_K_by_Kb3, 
 			      unique_session_keys]) 2);
@@ -317,25 +311,25 @@
 
 
 (*AUTHENTICATION OF B TO A*)
-Goal "[| Crypt K (Number Ta) : parts (spies evs);           \
+Goal "[| Crypt K (Number Ta) \\<in> parts (spies evs);           \
 \        Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}    \
-\        : parts (spies evs);                               \
+\        \\<in> parts (spies evs);                               \
 \        ~ Expired Ts evs;                                  \
-\        A ~: bad;  B ~: bad;  evs : kerberos_ban |]        \
-\     ==> Says B A (Crypt K (Number Ta)) : set evs";
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos_ban |]        \
+\     ==> Says B A (Crypt K (Number Ta)) \\<in> set evs";
 by (blast_tac (claset() addSDs [A_trusts_K_by_Kb2]
                         addSIs [lemma_B RS mp RS mp RS mp]
                         addSEs [Confidentiality_S RSN (2,rev_notE)]) 1);
 qed "Authentication_B";
 
 
-Goal "[| A ~: bad; B ~: bad; evs : kerberos_ban |]      ==>         \ 
-\        Key K ~: analz (spies evs) -->         \
+Goal "[| A \\<notin> bad; B \\<notin> bad; evs \\<in> kerberos_ban |]      ==>         \ 
+\        Key K \\<notin> analz (spies evs) -->         \
 \        Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})  \
-\        : set evs -->  \
-\         Crypt K {|Agent A, Number Ta|} : parts (spies evs) -->\
+\        \\<in> set evs -->  \
+\         Crypt K {|Agent A, Number Ta|} \\<in> parts (spies evs) -->\
 \        Says A B {|X, Crypt K {|Agent A, Number Ta|}|}  \
-\            : set evs";
+\            \\<in> set evs";
 by (etac kerberos_ban.induct 1);
 by (ftac Says_S_message_form 5 THEN assume_tac 5);     
 by (ftac Kb3_msg_in_parts_spies 5);
@@ -352,13 +346,13 @@
 
 
 (*AUTHENTICATION OF A TO B*)
-Goal "[| Crypt K {|Agent A, Number Ta|} : parts (spies evs);  \
+Goal "[| Crypt K {|Agent A, Number Ta|} \\<in> parts (spies evs);  \
 \        Crypt (shrK B) {|Number Ts, Agent A, Key K|}         \
-\        : parts (spies evs);                                 \
+\        \\<in> parts (spies evs);                                 \
 \        ~ Expired Ts evs;                                    \
-\        A ~: bad;  B ~: bad;  evs : kerberos_ban |]          \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> kerberos_ban |]          \
 \     ==> Says A B {|Crypt (shrK B) {|Number Ts, Agent A, Key K|}, \    
-\                    Crypt K {|Agent A, Number Ta|}|} : set evs";
+\                    Crypt K {|Agent A, Number Ta|}|} \\<in> set evs";
 by (blast_tac (claset() addSDs [B_trusts_K_by_Kb3]
                         addSIs [lemma_A RS mp RS mp RS mp]
                         addSEs [Confidentiality_S RSN (2,rev_notE)]) 1);
--- a/src/HOL/Auth/Kerberos_BAN.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -46,48 +46,48 @@
 inductive "kerberos_ban"
   intrs 
 
-    Nil  "[]: kerberos_ban"
+    Nil  "[] \\<in> kerberos_ban"
 
-    Fake "[| evs: kerberos_ban;  X: synth (analz (spies evs)) |]
-          ==> Says Spy B X # evs : kerberos_ban"
+    Fake "[| evsf \\<in> kerberos_ban;  X \\<in> synth (analz (spies evsf)) |]
+          ==> Says Spy B X # evsf \\<in> kerberos_ban"
 
 
-    Kb1  "[| evs1: kerberos_ban |]
+    Kb1  "[| evs1 \\<in> kerberos_ban |]
           ==> Says A Server {|Agent A, Agent B|} # evs1
-                :  kerberos_ban"
+                \\<in>  kerberos_ban"
 
 
-    Kb2  "[| evs2: kerberos_ban;  Key KAB ~: used evs2;
-             Says A' Server {|Agent A, Agent B|} : set evs2 |]
+    Kb2  "[| evs2 \\<in> kerberos_ban;  Key KAB \\<notin> used evs2;
+             Says A' Server {|Agent A, Agent B|} \\<in> set evs2 |]
           ==> Says Server A 
                 (Crypt (shrK A)
                    {|Number (CT evs2), Agent B, Key KAB,  
                     (Crypt (shrK B) {|Number (CT evs2), Agent A, Key KAB|})|}) 
-                # evs2 : kerberos_ban"
+                # evs2 \\<in> kerberos_ban"
 
 
-    Kb3  "[| evs3: kerberos_ban;  
+    Kb3  "[| evs3 \\<in> kerberos_ban;  
              Says S A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) 
-               : set evs3;
-             Says A Server {|Agent A, Agent B|} : set evs3;
+               \\<in> set evs3;
+             Says A Server {|Agent A, Agent B|} \\<in> set evs3;
              ~ Expired Ts evs3 |]
           ==> Says A B {|X, Crypt K {|Agent A, Number (CT evs3)|} |} 
-               # evs3 : kerberos_ban"
+               # evs3 \\<in> kerberos_ban"
 
 
-    Kb4  "[| evs4: kerberos_ban;  
+    Kb4  "[| evs4 \\<in> kerberos_ban;  
              Says A' B {|(Crypt (shrK B) {|Number Ts, Agent A, Key K|}), 
 		         (Crypt K {|Agent A, Number Ta|}) |}: set evs4;
              ~ Expired Ts evs4;  RecentAuth Ta evs4 |]
           ==> Says B A (Crypt K (Number Ta)) # evs4
-                : kerberos_ban"
+                \\<in> kerberos_ban"
 
          (*Old session keys may become compromised*)
-    Oops "[| evso: kerberos_ban;  
+    Oops "[| evso \\<in> kerberos_ban;  
              Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
-               : set evso;
+               \\<in> set evso;
              Expired Ts evso |]
-          ==> Notes Spy {|Number Ts, Key K|} # evso : kerberos_ban"
+          ==> Notes Spy {|Number Ts, Key K|} # evso \\<in> kerberos_ban"
 
 
 end
--- a/src/HOL/Auth/OtwayRees.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/OtwayRees.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -17,8 +17,8 @@
 
 
 (*A "possibility property": there are traces that reach the end*)
-Goal "[| B ~= Server |]   \
-\     ==> EX NA K. EX evs: otway.          \
+Goal "B ~= Server   \
+\     ==> \\<exists>NA K. \\<exists>evs \\<in> otway.          \
 \            Says B A {|Nonce NA, Crypt (shrK A) {|Nonce NA, Key K|}|} \
 \              : set evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
@@ -29,7 +29,7 @@
 by possibility_tac;
 result();
 
-Goal "[| Gets B X : set evs; evs : otway |] ==> EX A. Says A B X : set evs";
+Goal "[| Gets B X : set evs; evs : otway |] ==> \\<exists>A. Says A B X : set evs";
 by (etac rev_mp 1);
 by (etac otway.induct 1);
 by Auto_tac;
@@ -66,7 +66,7 @@
 bind_thm ("OR4_parts_knows_Spy",
           OR4_analz_knows_Spy RS (impOfSubs analz_subset_parts));
 
-(*For proving the easier theorems about X ~: parts (knows Spy evs).*)
+(*For proving the easier theorems about X \\<notin> parts (knows Spy evs).*)
 fun parts_induct_tac i = 
     etac otway.induct i			THEN 
     ftac Oops_parts_knows_Spy (i+7) THEN
@@ -75,7 +75,7 @@
     prove_simple_subgoals_tac  i;
 
 
-(** Theorems of the form X ~: parts (knows Spy evs) imply that NOBODY
+(** Theorems of the form X \\<notin> parts (knows Spy evs) imply that NOBODY
     sends messages containing X! **)
 
 (*Spy never sees a good agent's shared key!*)
@@ -94,25 +94,13 @@
 	Spy_analz_shrK RSN (2, rev_iffD1)];
 
 
-(*Nobody can have used non-existent keys!*)
-Goal "evs: otway ==> Key K ~: used evs --> K ~: keysFor(parts(knows Spy evs))";
-by (parts_induct_tac 1);
-(*Fake*)
-by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
-(*OR2, OR3*)
-by (ALLGOALS Blast_tac);
-qed_spec_mp "new_keys_not_used";
-Addsimps [new_keys_not_used];
-
-
-
 (*** Proofs involving analz ***)
 
 (*Describes the form of K and NA when the Server sends this message.  Also
   for Oops case.*)
 Goal "[| Says Server B {|NA, X, Crypt (shrK B) {|NB, Key K|}|} : set evs; \
 \        evs : otway |]                                           \
-\     ==> K ~: range shrK & (EX i. NA = Nonce i) & (EX j. NB = Nonce j)";
+\     ==> K \\<notin> range shrK & (\\<exists>i. NA = Nonce i) & (\\<exists>j. NB = Nonce j)";
 by (etac rev_mp 1);
 by (etac otway.induct 1);
 by (ALLGOALS Simp_tac);
@@ -154,7 +142,7 @@
 qed_spec_mp "analz_image_freshK";
 
 
-Goal "[| evs : otway;  KAB ~: range shrK |]               \
+Goal "[| evs : otway;  KAB \\<notin> range shrK |]               \
 \     ==> Key K : analz (insert (Key KAB) (knows Spy evs)) =  \
 \         (K = KAB | Key K : analz (knows Spy evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);
@@ -178,7 +166,7 @@
 (**** Authenticity properties relating to NA ****)
 
 (*Only OR1 can have caused such a part of a message to appear.*)
-Goal "[| A ~: bad;  evs : otway |]                             \
+Goal "[| A \\<notin> bad;  evs : otway |]                             \
 \ ==> Crypt (shrK A) {|NA, Agent A, Agent B|} : parts (knows Spy evs) --> \
 \     Says A B {|NA, Agent A, Agent B,                      \
 \                Crypt (shrK A) {|NA, Agent A, Agent B|}|}  \
@@ -189,7 +177,7 @@
 
 Goal "[| Gets B {|NA, Agent A, Agent B,                      \
 \                 Crypt (shrK A) {|NA, Agent A, Agent B|}|} : set evs; \
-\        A ~: bad; evs : otway |]                             \
+\        A \\<notin> bad; evs : otway |]                             \
 \      ==> Says A B {|NA, Agent A, Agent B,                      \
 \                     Crypt (shrK A) {|NA, Agent A, Agent B|}|}  \
 \            : set evs";
@@ -201,7 +189,7 @@
 
 Goal "[| Crypt (shrK A) {|NA, Agent A, Agent B|}: parts (knows Spy evs); \
 \        Crypt (shrK A) {|NA, Agent A, Agent C|}: parts (knows Spy evs); \
-\        evs : otway;  A ~: bad |]                                   \
+\        evs : otway;  A \\<notin> bad |]                                   \
 \     ==> B = C";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -214,10 +202,10 @@
 (*It is impossible to re-use a nonce in both OR1 and OR2.  This holds because
   OR2 encrypts Nonce NB.  It prevents the attack that can occur in the
   over-simplified version of this protocol: see OtwayRees_Bad.*)
-Goal "[| A ~: bad;  evs : otway |]                      \
+Goal "[| A \\<notin> bad;  evs : otway |]                      \
 \     ==> Crypt (shrK A) {|NA, Agent A, Agent B|} : parts (knows Spy evs) --> \
 \         Crypt (shrK A) {|NA', NA, Agent A', Agent A|}  \
-\           ~: parts (knows Spy evs)";
+\           \\<notin> parts (knows Spy evs)";
 by (parts_induct_tac 1);
 by Auto_tac;
 qed_spec_mp "no_nonce_OR1_OR2";
@@ -226,11 +214,11 @@
 
 (*Crucial property: If the encrypted message appears, and A has used NA
   to start a run, then it originated with the Server!*)
-Goal "[| A ~: bad;  evs : otway |]                                  \
+Goal "[| A \\<notin> bad;  evs : otway |]                                  \
 \     ==> Says A B {|NA, Agent A, Agent B,                          \
 \                    Crypt (shrK A) {|NA, Agent A, Agent B|}|} : set evs --> \
 \         Crypt (shrK A) {|NA, Key K|} : parts (knows Spy evs)          \
-\         --> (EX NB. Says Server B                                     \
+\         --> (\\<exists>NB. Says Server B                                     \
 \                        {|NA,                                          \
 \                          Crypt (shrK A) {|NA, Key K|},                \
 \                          Crypt (shrK B) {|NB, Key K|}|} : set evs)";
@@ -252,8 +240,8 @@
 Goal "[| Says A  B {|NA, Agent A, Agent B,                       \
 \                Crypt (shrK A) {|NA, Agent A, Agent B|}|} : set evs; \
 \        Gets A {|NA, Crypt (shrK A) {|NA, Key K|}|} : set evs; \
-\    A ~: bad;  evs : otway |]                              \
-\ ==> EX NB. Says Server B                                  \
+\    A \\<notin> bad;  evs : otway |]                              \
+\ ==> \\<exists>NB. Says Server B                                  \
 \              {|NA,                                        \
 \                Crypt (shrK A) {|NA, Key K|},              \
 \                Crypt (shrK B) {|NB, Key K|}|}             \
@@ -266,12 +254,12 @@
     Does not in itself guarantee security: an attack could violate 
     the premises, e.g. by having A=Spy **)
 
-Goal "[| A ~: bad;  B ~: bad;  evs : otway |]                  \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs : otway |]                  \
 \ ==> Says Server B                                            \
 \       {|NA, Crypt (shrK A) {|NA, Key K|},                    \
 \         Crypt (shrK B) {|NB, Key K|}|} : set evs -->         \
-\     Notes Spy {|NA, NB, Key K|} ~: set evs -->               \
-\     Key K ~: analz (knows Spy evs)";
+\     Notes Spy {|NA, NB, Key K|} \\<notin> set evs -->               \
+\     Key K \\<notin> analz (knows Spy evs)";
 by (etac otway.induct 1);
 by analz_knows_Spy_tac;
 by (ALLGOALS
@@ -291,9 +279,9 @@
 Goal "[| Says Server B                                           \
 \         {|NA, Crypt (shrK A) {|NA, Key K|},                    \
 \               Crypt (shrK B) {|NB, Key K|}|} : set evs;        \
-\        Notes Spy {|NA, NB, Key K|} ~: set evs;                 \
-\        A ~: bad;  B ~: bad;  evs : otway |]                    \
-\     ==> Key K ~: analz (knows Spy evs)";
+\        Notes Spy {|NA, NB, Key K|} \\<notin> set evs;                 \
+\        A \\<notin> bad;  B \\<notin> bad;  evs : otway |]                    \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addDs [Says_Server_message_form] addSEs [lemma]) 1);
 qed "Spy_not_see_encrypted_key";
 
@@ -303,9 +291,9 @@
 Goal "[| Says A  B {|NA, Agent A, Agent B,                       \
 \                    Crypt (shrK A) {|NA, Agent A, Agent B|}|} : set evs; \
 \        Gets A {|NA, Crypt (shrK A) {|NA, Key K|}|} : set evs; \
-\        ALL NB. Notes Spy {|NA, NB, Key K|} ~: set evs;         \
-\        A ~: bad;  B ~: bad;  evs : otway |]                    \
-\     ==> Key K ~: analz (knows Spy evs)";
+\        ALL NB. Notes Spy {|NA, NB, Key K|} \\<notin> set evs;         \
+\        A \\<notin> bad;  B \\<notin> bad;  evs : otway |]                    \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [A_trusts_OR4, Spy_not_see_encrypted_key]) 1);
 qed "A_gets_good_key";
 
@@ -316,8 +304,8 @@
   know anything about X: it does NOT have to have the right form.*)
 Goal "[| Crypt (shrK B) {|NA, NB, Agent A, Agent B|}       \
 \          : parts (knows Spy evs);  \
-\        B ~: bad;  evs : otway |]                         \
-\     ==> EX X. Says B Server                              \
+\        B \\<notin> bad;  evs : otway |]                         \
+\     ==> \\<exists>X. Says B Server                              \
 \                {|NA, Agent A, Agent B, X,                       \
 \                  Crypt (shrK B) {|NA, NB, Agent A, Agent B|}|}  \
 \                : set evs";
@@ -331,7 +319,7 @@
 
 Goal "[| Crypt (shrK B) {|NA, NB, Agent A, Agent B|} : parts(knows Spy evs); \
 \        Crypt (shrK B) {|NC, NB, Agent C, Agent B|} : parts(knows Spy evs); \
-\          evs : otway;  B ~: bad |]             \
+\          evs : otway;  B \\<notin> bad |]             \
 \        ==> NC = NA & C = A";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -342,7 +330,7 @@
 
 (*If the encrypted message appears, and B has used Nonce NB,
   then it originated with the Server!  Quite messy proof.*)
-Goal "[| B ~: bad;  evs : otway |]                                    \
+Goal "[| B \\<notin> bad;  evs : otway |]                                    \
 \ ==> Crypt (shrK B) {|NB, Key K|} : parts (knows Spy evs)            \
 \     --> (ALL X'. Says B Server                                      \
 \                    {|NA, Agent A, Agent B, X',                      \
@@ -371,7 +359,7 @@
 \                        Crypt (shrK B) {|NA, NB, Agent A, Agent B|} |} \
 \          : set evs;                                           \
 \        Gets B {|NA, X, Crypt (shrK B) {|NB, Key K|}|} : set evs;   \
-\        B ~: bad;  evs : otway |]                              \
+\        B \\<notin> bad;  evs : otway |]                              \
 \     ==> Says Server B                                         \
 \              {|NA,                                            \
 \                Crypt (shrK A) {|NA, Key K|},                  \
@@ -386,9 +374,9 @@
 \                        Crypt (shrK B) {|NA, NB, Agent A, Agent B|} |} \
 \          : set evs;                                           \
 \        Gets B {|NA, X, Crypt (shrK B) {|NB, Key K|}|} : set evs;   \
-\        Notes Spy {|NA, NB, Key K|} ~: set evs;                \
-\        A ~: bad;  B ~: bad;  evs : otway |]                   \
-\     ==> Key K ~: analz (knows Spy evs)";
+\        Notes Spy {|NA, NB, Key K|} \\<notin> set evs;                \
+\        A \\<notin> bad;  B \\<notin> bad;  evs : otway |]                   \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [B_trusts_OR3, Spy_not_see_encrypted_key]) 1);
 qed "B_gets_good_key";
 
@@ -396,8 +384,8 @@
 Goal "[| Says Server B                                       \
 \           {|NA, Crypt (shrK A) {|NA, Key K|},              \
 \             Crypt (shrK B) {|NB, Key K|}|} : set evs;      \
-\        B ~: bad;  evs : otway |]                           \
-\ ==> EX X. Says B Server {|NA, Agent A, Agent B, X,         \
+\        B \\<notin> bad;  evs : otway |]                           \
+\ ==> \\<exists>X. Says B Server {|NA, Agent A, Agent B, X,         \
 \                           Crypt (shrK B) {|NA, NB, Agent A, Agent B|} |} \
 \             : set evs";
 by (etac rev_mp 1);
@@ -414,8 +402,8 @@
 Goal "[| Gets A {|NA, Crypt (shrK A) {|NA, Key K|}|} : set evs;        \
 \        Says A  B {|NA, Agent A, Agent B,                                \
 \                    Crypt (shrK A) {|NA, Agent A, Agent B|}|} : set evs; \
-\        A ~: bad;  B ~: bad;  evs : otway |]                             \
-\ ==> EX NB X. Says B Server {|NA, Agent A, Agent B, X,               \
+\        A \\<notin> bad;  B \\<notin> bad;  evs : otway |]                             \
+\ ==> \\<exists>NB X. Says B Server {|NA, Agent A, Agent B, X,               \
 \                              Crypt (shrK B)  {|NA, NB, Agent A, Agent B|} |}\
 \                : set evs";
 by (blast_tac (claset() delrules [Gets_imp_knows_Spy RS parts.Inj]
--- a/src/HOL/Auth/OtwayRees.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/OtwayRees.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -13,29 +13,29 @@
 inductive "otway"
   intrs 
          (*Initial trace is empty*)
-    Nil  "[]: otway"
+    Nil  "[] \\<in> otway"
 
          (** These rules allow agents to send messages to themselves **)
 
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evsa: otway;  X: synth (analz (knows Spy evsa)) |]
-          ==> Says Spy B X  # evsa : otway"
+    Fake "[| evsf \\<in> otway;  X \\<in> synth (analz (knows Spy evsf)) |]
+          ==> Says Spy B X  # evsf : otway"
 
          (*A message that has been sent can be received by the
            intended recipient.*)
-    Reception "[| evsr: otway;  Says A B X : set evsr |]
+    Reception "[| evsr \\<in> otway;  Says A B X : set evsr |]
                ==> Gets B X # evsr : otway"
 
          (*Alice initiates a protocol run*)
-    OR1  "[| evs1: otway;  Nonce NA ~: used evs1 |]
+    OR1  "[| evs1 \\<in> otway;  Nonce NA \\<notin> used evs1 |]
           ==> Says A B {|Nonce NA, Agent A, Agent B, 
                          Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |} 
                  # evs1 : otway"
 
          (*Bob's response to Alice's message.  Note that NB is encrypted.*)
-    OR2  "[| evs2: otway;  Nonce NB ~: used evs2;
+    OR2  "[| evs2 \\<in> otway;  Nonce NB \\<notin> used evs2;
              Gets B {|Nonce NA, Agent A, Agent B, X|} : set evs2 |]
           ==> Says B Server 
                   {|Nonce NA, Agent A, Agent B, X, 
@@ -46,7 +46,7 @@
          (*The Server receives Bob's message and checks that the three NAs
            match.  Then he sends a new session key to Bob with a packet for
            forwarding to Alice.*)
-    OR3  "[| evs3: otway;  Key KAB ~: used evs3;
+    OR3  "[| evs3 \\<in> otway;  Key KAB \\<notin> used evs3;
              Gets Server 
                   {|Nonce NA, Agent A, Agent B, 
                     Crypt (shrK A) {|Nonce NA, Agent A, Agent B|}, 
@@ -61,7 +61,7 @@
          (*Bob receives the Server's (?) message and compares the Nonces with
 	   those in the message he previously sent the Server.
            Need B ~= Server because we allow messages to self.*)
-    OR4  "[| evs4: otway;  B ~= Server;
+    OR4  "[| evs4 \\<in> otway;  B ~= Server;
              Says B Server {|Nonce NA, Agent A, Agent B, X', 
                              Crypt (shrK B)
                                    {|Nonce NA, Nonce NB, Agent A, Agent B|}|}
@@ -72,7 +72,7 @@
 
          (*This message models possible leaks of session keys.  The nonces
            identify the protocol run.*)
-    Oops "[| evso: otway;  
+    Oops "[| evso \\<in> otway;  
              Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
                : set evso |]
           ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso : otway"
--- a/src/HOL/Auth/OtwayRees_AN.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/OtwayRees_AN.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -17,10 +17,10 @@
 
 
 (*A "possibility property": there are traces that reach the end*)
-Goal "[| B ~= Server |]   \
-\     ==> EX K. EX NA. EX evs: otway.                                      \
+Goal "B ~= Server   \
+\     ==> \\<exists>K. \\<exists>NA. \\<exists>evs \\<in> otway.                                      \
 \          Says B A (Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|}) \
-\            : set evs";
+\            \\<in> set evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (otway.Nil RS 
           otway.OR1 RS otway.Reception RS
@@ -29,14 +29,14 @@
 by possibility_tac;
 result();
 
-Goal "[| Gets B X : set evs; evs : otway |] ==> EX A. Says A B X : set evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> otway |] ==> \\<exists>A. Says A B X \\<in> set evs";
 by (etac rev_mp 1);
 by (etac otway.induct 1);
 by Auto_tac;
 qed"Gets_imp_Says";
 
 (*Must be proved separately for each protocol*)
-Goal "[| Gets B X : set evs; evs : otway |]  ==> X : knows Spy evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> otway |]  ==> X \\<in> knows Spy evs";
 by (blast_tac (claset() addSDs [Gets_imp_Says, Says_imp_knows_Spy]) 1);
 qed"Gets_imp_knows_Spy";
 AddDs [Gets_imp_knows_Spy RS parts.Inj];
@@ -46,20 +46,20 @@
 
 (** For reasoning about the encrypted portion of messages **)
 
-Goal "[| Gets B {|X, Crypt(shrK B) X'|} : set evs;  evs : otway |] ==> \
-\          X : analz (knows Spy evs)";
+Goal "[| Gets B {|X, Crypt(shrK B) X'|} \\<in> set evs;  evs \\<in> otway |]  \
+\     ==> X \\<in> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [Gets_imp_knows_Spy RS analz.Inj]) 1);
 qed "OR4_analz_knows_Spy";
 
-Goal "Says Server B {|X, Crypt K' {|NB, a, Agent B, K|}|} : set evs \
-\     ==> K : parts (knows Spy evs)";
+Goal "Says Server B {|X, Crypt K' {|NB, a, Agent B, K|}|} \\<in> set evs \
+\     ==> K \\<in> parts (knows Spy evs)";
 by (Blast_tac 1);
 qed "Oops_parts_knows_Spy";
 
 bind_thm ("OR4_parts_knows_Spy",
           OR4_analz_knows_Spy RS (impOfSubs analz_subset_parts));
 
-(*For proving the easier theorems about X ~: parts (knows Spy evs).*)
+(*For proving the easier theorems about X \\<notin> parts (knows Spy evs).*)
 fun parts_induct_tac i = 
     etac otway.induct i			THEN 
     ftac Oops_parts_knows_Spy (i+7) THEN
@@ -67,17 +67,17 @@
     prove_simple_subgoals_tac  i;
 
 
-(** Theorems of the form X ~: parts (knows Spy evs) imply that NOBODY
+(** Theorems of the form X \\<notin> parts (knows Spy evs) imply that NOBODY
     sends messages containing X! **)
 
 (*Spy never sees a good agent's shared key!*)
-Goal "evs : otway ==> (Key (shrK A) : parts (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> otway ==> (Key (shrK A) \\<in> parts (knows Spy evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
-Goal "evs : otway ==> (Key (shrK A) : analz (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> otway ==> (Key (shrK A) \\<in> analz (knows Spy evs)) = (A \\<in> bad)";
 by (auto_tac(claset() addDs [impOfSubs analz_subset_parts], simpset()));
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
@@ -86,32 +86,15 @@
 	Spy_analz_shrK RSN (2, rev_iffD1)];
 
 
-(*Nobody can have used non-existent keys!*)
-Goal "evs : otway ==> Key K ~: used evs --> K ~: keysFor (parts (knows Spy evs))";
-by (parts_induct_tac 1);
-(*Fake*)
-by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
-(*OR3*)
-by (Blast_tac 1);
-qed_spec_mp "new_keys_not_used";
-
-bind_thm ("new_keys_not_analzd",
-          [analz_subset_parts RS keysFor_mono,
-           new_keys_not_used] MRS contra_subsetD);
-
-Addsimps [new_keys_not_used, new_keys_not_analzd];
-
-
-
 (*** Proofs involving analz ***)
 
 (*Describes the form of K and NA when the Server sends this message.*)
 Goal "[| Says Server B                                           \
 \           {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},    \
 \             Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}   \
-\          : set evs;                                            \
-\        evs : otway |]                                          \
-\     ==> K ~: range shrK & (EX i. NA = Nonce i) & (EX j. NB = Nonce j)";
+\          \\<in> set evs;                                            \
+\        evs \\<in> otway |]                                          \
+\     ==> K \\<notin> range shrK & (\\<exists>i. NA = Nonce i) & (\\<exists>j. NB = Nonce j)";
 by (etac rev_mp 1);
 by (etac otway.induct 1);
 by (ALLGOALS Asm_simp_tac);
@@ -139,10 +122,10 @@
 (** Session keys are not used to encrypt other session keys **)
 
 (*The equality makes the induction hypothesis easier to apply*)
-Goal "evs : otway ==>                                 \
+Goal "evs \\<in> otway ==>                                 \
 \  ALL K KK. KK <= -(range shrK) -->                  \
-\         (Key K : analz (Key`KK Un (knows Spy evs))) =  \
-\         (K : KK | Key K : analz (knows Spy evs))";
+\         (Key K \\<in> analz (Key`KK Un (knows Spy evs))) =  \
+\         (K \\<in> KK | Key K \\<in> analz (knows Spy evs))";
 by (etac otway.induct 1);
 by analz_knows_Spy_tac;
 by (REPEAT_FIRST (resolve_tac [allI, impI]));
@@ -153,9 +136,9 @@
 qed_spec_mp "analz_image_freshK";
 
 
-Goal "[| evs : otway;  KAB ~: range shrK |] ==>       \
-\     Key K : analz (insert (Key KAB) (knows Spy evs)) =  \
-\     (K = KAB | Key K : analz (knows Spy evs))";
+Goal "[| evs \\<in> otway;  KAB \\<notin> range shrK |] ==>       \
+\     Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) =  \
+\     (K = KAB | Key K \\<in> analz (knows Spy evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);
 qed "analz_insert_freshK";
 
@@ -165,12 +148,12 @@
 Goal "[| Says Server B                                           \
 \         {|Crypt (shrK A) {|NA, Agent A, Agent B, K|},         \
 \           Crypt (shrK B) {|NB, Agent A, Agent B, K|}|}        \
-\        : set evs;                                             \
+\        \\<in> set evs;                                             \
 \       Says Server B'                                          \
 \         {|Crypt (shrK A') {|NA', Agent A', Agent B', K|},     \
 \           Crypt (shrK B') {|NB', Agent A', Agent B', K|}|}    \
-\        : set evs;                                             \
-\       evs : otway |]                                          \
+\        \\<in> set evs;                                             \
+\       evs \\<in> otway |]                                          \
 \    ==> A=A' & B=B' & NA=NA' & NB=NB'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -185,12 +168,12 @@
 (**** Authenticity properties relating to NA ****)
 
 (*If the encrypted message appears then it originated with the Server!*)
-Goal "[| A ~: bad;  A ~= B;  evs : otway |]                 \
-\     ==> Crypt (shrK A) {|NA, Agent A, Agent B, Key K|} : parts (knows Spy evs) \
-\      --> (EX NB. Says Server B                                          \
+Goal "[| A \\<notin> bad;  A ~= B;  evs \\<in> otway |]                 \
+\     ==> Crypt (shrK A) {|NA, Agent A, Agent B, Key K|} \\<in> parts (knows Spy evs) \
+\      --> (\\<exists>NB. Says Server B                                          \
 \                   {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},     \
 \                     Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}    \
-\                   : set evs)";
+\                   \\<in> set evs)";
 by (parts_induct_tac 1);
 by (Blast_tac 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [ex_disj_distrib])));
@@ -202,12 +185,12 @@
 (*Corollary: if A receives B's OR4 message then it originated with the Server.
   Freshness may be inferred from nonce NA.*)
 Goal "[| Gets A (Crypt (shrK A) {|NA, Agent A, Agent B, Key K|})  \
-\         : set evs;                                                 \
-\        A ~: bad;  A ~= B;  evs : otway |]                          \
-\     ==> EX NB. Says Server B                                       \
+\         \\<in> set evs;                                                 \
+\        A \\<notin> bad;  A ~= B;  evs \\<in> otway |]                          \
+\     ==> \\<exists>NB. Says Server B                                       \
 \                 {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},  \
 \                   Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} \
-\                : set evs";
+\                \\<in> set evs";
 by (blast_tac (claset() addSIs [NA_Crypt_imp_Server_msg]) 1);
 qed "A_trusts_OR4";
 
@@ -216,13 +199,13 @@
     Does not in itself guarantee security: an attack could violate 
     the premises, e.g. by having A=Spy **)
 
-Goal "[| A ~: bad;  B ~: bad;  evs : otway |]                   \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs \\<in> otway |]                   \
 \     ==> Says Server B                                         \
 \          {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},    \
 \            Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}   \
-\         : set evs -->                                         \
-\         Notes Spy {|NA, NB, Key K|} ~: set evs -->            \
-\         Key K ~: analz (knows Spy evs)";
+\         \\<in> set evs -->                                         \
+\         Notes Spy {|NA, NB, Key K|} \\<notin> set evs -->            \
+\         Key K \\<notin> analz (knows Spy evs)";
 by (etac otway.induct 1);
 by analz_knows_Spy_tac;
 by (ALLGOALS
@@ -242,10 +225,10 @@
 Goal "[| Says Server B                                           \
 \           {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},    \
 \             Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}   \
-\          : set evs;                                            \
-\        Notes Spy {|NA, NB, Key K|} ~: set evs;                 \
-\        A ~: bad;  B ~: bad;  evs : otway |]                    \
-\     ==> Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs;                                            \
+\        Notes Spy {|NA, NB, Key K|} \\<notin> set evs;                 \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> otway |]                    \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (ftac Says_Server_message_form 1 THEN assume_tac 1);
 by (blast_tac (claset() addSEs [lemma]) 1);
 qed "Spy_not_see_encrypted_key";
@@ -254,10 +237,10 @@
 (*A's guarantee.  The Oops premise quantifies over NB because A cannot know
   what it is.*)
 Goal "[| Gets A (Crypt (shrK A) {|NA, Agent A, Agent B, Key K|})  \
-\         : set evs;                                                 \
-\        ALL NB. Notes Spy {|NA, NB, Key K|} ~: set evs;             \
-\        A ~: bad;  B ~: bad;  A ~= B;  evs : otway |]               \
-\     ==> Key K ~: analz (knows Spy evs)";
+\         \\<in> set evs;                                                 \
+\        ALL NB. Notes Spy {|NA, NB, Key K|} \\<notin> set evs;             \
+\        A \\<notin> bad;  B \\<notin> bad;  A ~= B;  evs \\<in> otway |]               \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [A_trusts_OR4, Spy_not_see_encrypted_key]) 1);
 qed "A_gets_good_key";
 
@@ -265,12 +248,12 @@
 (**** Authenticity properties relating to NB ****)
 
 (*If the encrypted message appears then it originated with the Server!*)
-Goal "[| B ~: bad;  A ~= B;  evs : otway |]                              \
-\ ==> Crypt (shrK B) {|NB, Agent A, Agent B, Key K|} : parts (knows Spy evs) \
-\     --> (EX NA. Says Server B                                          \
+Goal "[| B \\<notin> bad;  A ~= B;  evs \\<in> otway |]                              \
+\ ==> Crypt (shrK B) {|NB, Agent A, Agent B, Key K|} \\<in> parts (knows Spy evs) \
+\     --> (\\<exists>NA. Says Server B                                          \
 \                  {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},     \
 \                    Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}    \
-\                  : set evs)";
+\                  \\<in> set evs)";
 by (parts_induct_tac 1);
 by (Blast_tac 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [ex_disj_distrib])));
@@ -282,21 +265,21 @@
 (*Guarantee for B: if it gets a well-formed certificate then the Server
   has sent the correct message in round 3.*)
 Goal "[| Gets B {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} \
-\          : set evs;                                                    \
-\        B ~: bad;  A ~= B;  evs : otway |]                              \
-\     ==> EX NA. Says Server B                                           \
+\          \\<in> set evs;                                                    \
+\        B \\<notin> bad;  A ~= B;  evs \\<in> otway |]                              \
+\     ==> \\<exists>NA. Says Server B                                           \
 \                  {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},     \
 \                    Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}    \
-\                  : set evs";
+\                  \\<in> set evs";
 by (blast_tac (claset() addSIs [NB_Crypt_imp_Server_msg]) 1);
 qed "B_trusts_OR3";
 
 
 (*The obvious combination of B_trusts_OR3 with Spy_not_see_encrypted_key*)
 Goal "[| Gets B {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|} \
-\         : set evs;                                                     \
-\        ALL NA. Notes Spy {|NA, NB, Key K|} ~: set evs;                 \
-\        A ~: bad;  B ~: bad;  A ~= B;  evs : otway |]                   \
-\     ==> Key K ~: analz (knows Spy evs)";
+\         \\<in> set evs;                                                     \
+\        ALL NA. Notes Spy {|NA, NB, Key K|} \\<notin> set evs;                 \
+\        A \\<notin> bad;  B \\<notin> bad;  A ~= B;  evs \\<in> otway |]                   \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addDs [B_trusts_OR3, Spy_not_see_encrypted_key]) 1);
 qed "B_gets_good_key";
--- a/src/HOL/Auth/OtwayRees_AN.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/OtwayRees_AN.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -28,50 +28,50 @@
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evs: otway;  X: synth (analz (knows Spy evs)) |]
-          ==> Says Spy B X  # evs : otway"
+    Fake "[| evs \\<in> otway;  X \\<in> synth (analz (knows Spy evs)) |]
+          ==> Says Spy B X  # evs \\<in> otway"
 
          (*A message that has been sent can be received by the
            intended recipient.*)
-    Reception "[| evsr: otway;  Says A B X : set evsr |]
-               ==> Gets B X # evsr : otway"
+    Reception "[| evsr \\<in> otway;  Says A B X \\<in>set evsr |]
+               ==> Gets B X # evsr \\<in> otway"
 
          (*Alice initiates a protocol run*)
-    OR1  "[| evs1: otway |]
-          ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 : otway"
+    OR1  "[| evs1 \\<in> otway |]
+          ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \\<in> otway"
 
          (*Bob's response to Alice's message.*)
-    OR2  "[| evs2: otway;  
-             Gets B {|Agent A, Agent B, Nonce NA|} : set evs2 |]
+    OR2  "[| evs2 \\<in> otway;  
+             Gets B {|Agent A, Agent B, Nonce NA|} \\<in>set evs2 |]
           ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
-                 # evs2 : otway"
+                 # evs2 \\<in> otway"
 
          (*The Server receives Bob's message.  Then he sends a new
            session key to Bob with a packet for forwarding to Alice.*)
-    OR3  "[| evs3: otway;  Key KAB ~: used evs3;
+    OR3  "[| evs3 \\<in> otway;  Key KAB \\<notin> used evs3;
              Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
-               : set evs3 |]
+               \\<in>set evs3 |]
           ==> Says Server B 
                {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key KAB|},
                  Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key KAB|}|}
-              # evs3 : otway"
+              # evs3 \\<in> otway"
 
          (*Bob receives the Server's (?) message and compares the Nonces with
 	   those in the message he previously sent the Server.
            Need B ~= Server because we allow messages to self.*)
-    OR4  "[| evs4: otway;  B ~= Server; 
-             Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} : set evs4;
+    OR4  "[| evs4 \\<in> otway;  B ~= Server; 
+             Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \\<in>set evs4;
              Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|}
-               : set evs4 |]
-          ==> Says B A X # evs4 : otway"
+               \\<in>set evs4 |]
+          ==> Says B A X # evs4 \\<in> otway"
 
          (*This message models possible leaks of session keys.  The nonces
            identify the protocol run.  B is not assumed to know shrK A.*)
-    Oops "[| evso: otway;  
+    Oops "[| evso \\<in> otway;  
              Says Server B 
                       {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|}, 
                         Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|}
-               : set evso |]
-          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso : otway"
+               \\<in>set evso |]
+          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \\<in> otway"
 
 end
--- a/src/HOL/Auth/OtwayRees_Bad.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/OtwayRees_Bad.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -19,10 +19,10 @@
 AddDs  [impOfSubs analz_subset_parts, impOfSubs Fake_parts_insert];
 
 (*A "possibility property": there are traces that reach the end*)
-Goal "[| A ~= B; B ~= Server |]   \
-\     ==> EX K. EX NA. EX evs: otway.          \
+Goal "B ~= Server   \
+\     ==> \\<exists>K. \\<exists>NA. \\<exists>evs \\<in> otway.          \
 \           Says B A {|Nonce NA, Crypt (shrK A) {|Nonce NA, Key K|}|} \
-\             : set evs";
+\             \\<in> set evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (otway.Nil RS 
           otway.OR1 RS otway.Reception RS
@@ -31,14 +31,14 @@
 by possibility_tac;
 result();
 
-Goal "[| Gets B X : set evs; evs : otway |] ==> EX A. Says A B X : set evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> otway |] ==> \\<exists>A. Says A B X \\<in> set evs";
 by (etac rev_mp 1);
 by (etac otway.induct 1);
 by Auto_tac;
 qed"Gets_imp_Says";
 
 (*Must be proved separately for each protocol*)
-Goal "[| Gets B X : set evs; evs : otway |]  ==> X : knows Spy evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> otway |]  ==> X \\<in> knows Spy evs";
 by (blast_tac (claset() addSDs [Gets_imp_Says, Says_imp_knows_Spy]) 1);
 qed"Gets_imp_knows_Spy";
 AddDs [Gets_imp_knows_Spy RS parts.Inj];
@@ -49,18 +49,18 @@
 
 (** For reasoning about the encrypted portion of messages **)
 
-Goal "[| Gets B {|N, Agent A, Agent B, X|} : set evs;  evs : otway |] \
-\     ==> X : analz (knows Spy evs)";
+Goal "[| Gets B {|N, Agent A, Agent B, X|} \\<in> set evs;  evs \\<in> otway |] \
+\     ==> X \\<in> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [Gets_imp_knows_Spy RS analz.Inj]) 1);
 qed "OR2_analz_knows_Spy";
 
-Goal "[| Gets B {|N, X, Crypt (shrK B) X'|} : set evs;  evs : otway |] \
-\     ==> X : analz (knows Spy evs)";
+Goal "[| Gets B {|N, X, Crypt (shrK B) X'|} \\<in> set evs;  evs \\<in> otway |] \
+\     ==> X \\<in> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [Gets_imp_knows_Spy RS analz.Inj]) 1);
 qed "OR4_analz_knows_Spy";
 
-Goal "Says Server B {|NA, X, Crypt K' {|NB,K|}|} : set evs \
-\     ==> K : parts (knows Spy evs)";
+Goal "Says Server B {|NA, X, Crypt K' {|NB,K|}|} \\<in> set evs \
+\     ==> K \\<in> parts (knows Spy evs)";
 by (Blast_tac 1);
 qed "Oops_parts_knows_Spy";
 
@@ -69,7 +69,7 @@
 bind_thm ("OR4_parts_knows_Spy",
           OR4_analz_knows_Spy RS (impOfSubs analz_subset_parts));
 
-(*For proving the easier theorems about X ~: parts (knows Spy evs).*)
+(*For proving the easier theorems about X \\<notin> parts (knows Spy evs).*)
 fun parts_induct_tac i = 
     etac otway.induct i			THEN 
     ftac Oops_parts_knows_Spy (i+7) THEN
@@ -78,17 +78,17 @@
     prove_simple_subgoals_tac  i;
 
 
-(** Theorems of the form X ~: parts (knows Spy evs) imply that NOBODY
+(** Theorems of the form X \\<notin> parts (knows Spy evs) imply that NOBODY
     sends messages containing X! **)
 
 (*Spy never sees a good agent's shared key!*)
-Goal "evs : otway ==> (Key (shrK A) : parts (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> otway ==> (Key (shrK A) \\<in> parts (knows Spy evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
-Goal "evs : otway ==> (Key (shrK A) : analz (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> otway ==> (Key (shrK A) \\<in> analz (knows Spy evs)) = (A \\<in> bad)";
 by (auto_tac(claset() addDs [impOfSubs analz_subset_parts], simpset()));
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
@@ -97,25 +97,13 @@
 	Spy_analz_shrK RSN (2, rev_iffD1)];
 
 
-(*Nobody can have used non-existent keys!*)
-Goal "evs : otway ==> Key K ~: used evs --> K ~: keysFor (parts (knows Spy evs))";
-by (parts_induct_tac 1);
-(*Fake*)
-by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
-(*OR2, OR3*)
-by (ALLGOALS Blast_tac);
-qed_spec_mp "new_keys_not_used";
-Addsimps [new_keys_not_used];
-
-
-
 (*** Proofs involving analz ***)
 
 (*Describes the form of K and NA when the Server sends this message.  Also
   for Oops case.*)
-Goal "[| Says Server B {|NA, X, Crypt (shrK B) {|NB, Key K|}|} : set evs; \
-\        evs : otway |]                                           \
-\  ==> K ~: range shrK & (EX i. NA = Nonce i) & (EX j. NB = Nonce j)";
+Goal "[| Says Server B {|NA, X, Crypt (shrK B) {|NB, Key K|}|} \\<in> set evs; \
+\        evs \\<in> otway |]                                           \
+\  ==> K \\<notin> range shrK & (\\<exists>i. NA = Nonce i) & (\\<exists>j. NB = Nonce j)";
 by (etac rev_mp 1);
 by (etac otway.induct 1);
 by (ALLGOALS Simp_tac);
@@ -134,8 +122,8 @@
 (****
  The following is to prove theorems of the form
 
-  Key K : analz (insert (Key KAB) (knows Spy evs)) ==>
-  Key K : analz (knows Spy evs)
+  Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) ==>
+  Key K \\<in> analz (knows Spy evs)
 
  A more general formula must be proved inductively.
 ****)
@@ -144,10 +132,10 @@
 (** Session keys are not used to encrypt other session keys **)
 
 (*The equality makes the induction hypothesis easier to apply*)
-Goal "evs : otway ==>                                 \
-\  ALL K KK. KK <= - (range shrK) -->                 \
-\         (Key K : analz (Key`KK Un (knows Spy evs))) =  \
-\         (K : KK | Key K : analz (knows Spy evs))";
+Goal "evs \\<in> otway ==>                                 \
+\  \\<forall>K KK. KK <= - (range shrK) -->                 \
+\         (Key K \\<in> analz (Key`KK Un (knows Spy evs))) =  \
+\         (K \\<in> KK | Key K \\<in> analz (knows Spy evs))";
 by (etac otway.induct 1);
 by analz_knows_Spy_tac;
 by (REPEAT_FIRST (resolve_tac [allI, impI]));
@@ -158,18 +146,18 @@
 qed_spec_mp "analz_image_freshK";
 
 
-Goal "[| evs : otway;  KAB ~: range shrK |] ==>       \
-\     Key K : analz (insert (Key KAB) (knows Spy evs)) =  \
-\     (K = KAB | Key K : analz (knows Spy evs))";
+Goal "[| evs \\<in> otway;  KAB \\<notin> range shrK |] ==>       \
+\     Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) =  \
+\     (K = KAB | Key K \\<in> analz (knows Spy evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);
 qed "analz_insert_freshK";
 
 
 (*** The Key K uniquely identifies the Server's  message. **)
 
-Goal "[| Says Server B {|NA, X, Crypt (shrK B) {|NB, K|}|}   : set evs; \ 
-\        Says Server B' {|NA',X',Crypt (shrK B') {|NB',K|}|} : set evs; \
-\        evs : otway |] ==> X=X' & B=B' & NA=NA' & NB=NB'";
+Goal "[| Says Server B {|NA, X, Crypt (shrK B) {|NB, K|}|}   \\<in> set evs; \ 
+\        Says Server B' {|NA',X',Crypt (shrK B') {|NB',K|}|} \\<in> set evs; \
+\        evs \\<in> otway |] ==> X=X' & B=B' & NA=NA' & NB=NB'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (etac otway.induct 1);
@@ -183,12 +171,12 @@
     Does not in itself guarantee security: an attack could violate 
     the premises, e.g. by having A=Spy **)
 
-Goal "[| A ~: bad;  B ~: bad;  evs : otway |]                      \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs \\<in> otway |]                      \
 \     ==> Says Server B                                            \
 \           {|NA, Crypt (shrK A) {|NA, Key K|},                    \
-\             Crypt (shrK B) {|NB, Key K|}|} : set evs -->         \
-\         Notes Spy {|NA, NB, Key K|} ~: set evs -->               \
-\         Key K ~: analz (knows Spy evs)";
+\             Crypt (shrK B) {|NB, Key K|}|} \\<in> set evs -->         \
+\         Notes Spy {|NA, NB, Key K|} \\<notin> set evs -->               \
+\         Key K \\<notin> analz (knows Spy evs)";
 by (etac otway.induct 1);
 by analz_knows_Spy_tac;
 by (ALLGOALS
@@ -207,10 +195,10 @@
 
 Goal "[| Says Server B                                           \
 \         {|NA, Crypt (shrK A) {|NA, Key K|},                    \
-\               Crypt (shrK B) {|NB, Key K|}|} : set evs;        \
-\        Notes Spy {|NA, NB, Key K|} ~: set evs;                 \
-\        A ~: bad;  B ~: bad;  evs : otway |]                    \
-\     ==> Key K ~: analz (knows Spy evs)";
+\               Crypt (shrK B) {|NB, Key K|}|} \\<in> set evs;        \
+\        Notes Spy {|NA, NB, Key K|} \\<notin> set evs;                 \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> otway |]                    \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (ftac Says_Server_message_form 1 THEN assume_tac 1);
 by (blast_tac (claset() addSEs [lemma]) 1);
 qed "Spy_not_see_encrypted_key";
@@ -221,10 +209,10 @@
 (*Only OR1 can have caused such a part of a message to appear.
   The premise A ~= B prevents OR2's similar-looking cryptogram from being
   picked up.  Original Otway-Rees doesn't need it.*)
-Goal "[| A ~: bad;  A ~= B;  evs : otway |]                \
-\     ==> Crypt (shrK A) {|NA, Agent A, Agent B|} : parts (knows Spy evs) --> \
+Goal "[| A \\<notin> bad;  A ~= B;  evs \\<in> otway |]                \
+\     ==> Crypt (shrK A) {|NA, Agent A, Agent B|} \\<in> parts (knows Spy evs) --> \
 \         Says A B {|NA, Agent A, Agent B,                  \
-\                    Crypt (shrK A) {|NA, Agent A, Agent B|}|}  : set evs";
+\                    Crypt (shrK A) {|NA, Agent A, Agent B|}|}  \\<in> set evs";
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
 qed_spec_mp "Crypt_imp_OR1";
@@ -235,15 +223,15 @@
   The premise A ~= B allows use of Crypt_imp_OR1*)
 (*Only it is FALSE.  Somebody could make a fake message to Server
           substituting some other nonce NA' for NB.*)
-Goal "[| A ~: bad;  A ~= B;  evs : otway |]                                \
-\     ==> Crypt (shrK A) {|NA, Key K|} : parts (knows Spy evs) -->    \
+Goal "[| A \\<notin> bad;  A ~= B;  evs \\<in> otway |]                                \
+\     ==> Crypt (shrK A) {|NA, Key K|} \\<in> parts (knows Spy evs) -->    \
 \         Says A B {|NA, Agent A, Agent B,                        \
 \                    Crypt (shrK A) {|NA, Agent A, Agent B|}|}    \
-\          : set evs -->                                          \
-\         (EX B NB. Says Server B                                 \
+\          \\<in> set evs -->                                          \
+\         (\\<exists>B NB. Says Server B                                 \
 \              {|NA,                                              \
 \                Crypt (shrK A) {|NA, Key K|},                    \
-\                Crypt (shrK B) {|NB, Key K|}|}  : set evs)";
+\                Crypt (shrK B) {|NB, Key K|}|}  \\<in> set evs)";
 by (parts_induct_tac 1);
 (*Fake*)
 by (Blast_tac 1);
@@ -261,11 +249,11 @@
            {|Nonce NA, Agent Aa, Agent A,
              Crypt (shrK Aa) {|Nonce NA, Agent Aa, Agent A|}, Nonce NB,
              Crypt (shrK A) {|Nonce NA, Agent Aa, Agent A|}|}
-          : set evs3;
+          \\<in> set evs3;
           Says A B
            {|Nonce NB, Agent A, Agent B,
              Crypt (shrK A) {|Nonce NB, Agent A, Agent B|}|}
-          : set evs3;
+          \\<in> set evs3;
 *)
 writeln "GIVE UP! on NA_Crypt_imp_Server_msg";
 
--- a/src/HOL/Auth/OtwayRees_Bad.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/OtwayRees_Bad.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -17,66 +17,66 @@
 inductive otway
   intrs 
          (*Initial trace is empty*)
-    Nil  "[]: otway"
+    Nil  "[] \\<in> otway"
 
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evs: otway;  X: synth (analz (knows Spy evs)) |]
-          ==> Says Spy B X  # evs : otway"
+    Fake "[| evsf \\<in> otway;  X \\<in> synth (analz (knows Spy evsf)) |]
+          ==> Says Spy B X  # evsf \\<in> otway"
 
          (*A message that has been sent can be received by the
            intended recipient.*)
-    Reception "[| evsr: otway;  Says A B X : set evsr |]
-               ==> Gets B X # evsr : otway"
+    Reception "[| evsr \\<in> otway;  Says A B X \\<in> set evsr |]
+               ==> Gets B X # evsr \\<in> otway"
 
          (*Alice initiates a protocol run*)
-    OR1  "[| evs1: otway;  Nonce NA ~: used evs1 |]
+    OR1  "[| evs1 \\<in> otway;  Nonce NA \\<notin> used evs1 |]
           ==> Says A B {|Nonce NA, Agent A, Agent B, 
                          Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |} 
-                 # evs1 : otway"
+                 # evs1 \\<in> otway"
 
          (*Bob's response to Alice's message. 
            This variant of the protocol does NOT encrypt NB.*)
-    OR2  "[| evs2: otway;  Nonce NB ~: used evs2;
-             Gets B {|Nonce NA, Agent A, Agent B, X|} : set evs2 |]
+    OR2  "[| evs2 \\<in> otway;  Nonce NB \\<notin> used evs2;
+             Gets B {|Nonce NA, Agent A, Agent B, X|} \\<in> set evs2 |]
           ==> Says B Server 
                   {|Nonce NA, Agent A, Agent B, X, Nonce NB,
                     Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
-                 # evs2 : otway"
+                 # evs2 \\<in> otway"
 
          (*The Server receives Bob's message and checks that the three NAs
            match.  Then he sends a new session key to Bob with a packet for
            forwarding to Alice.*)
-    OR3  "[| evs3: otway;  Key KAB ~: used evs3;
+    OR3  "[| evs3 \\<in> otway;  Key KAB \\<notin> used evs3;
              Gets Server 
                   {|Nonce NA, Agent A, Agent B, 
                     Crypt (shrK A) {|Nonce NA, Agent A, Agent B|}, 
                     Nonce NB, 
                     Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
-               : set evs3 |]
+               \\<in> set evs3 |]
           ==> Says Server B 
                   {|Nonce NA, 
                     Crypt (shrK A) {|Nonce NA, Key KAB|},
                     Crypt (shrK B) {|Nonce NB, Key KAB|}|}
-                 # evs3 : otway"
+                 # evs3 \\<in> otway"
 
          (*Bob receives the Server's (?) message and compares the Nonces with
 	   those in the message he previously sent the Server.
            Need B ~= Server because we allow messages to self.*)
-    OR4  "[| evs4: otway;  B ~= Server;
+    OR4  "[| evs4 \\<in> otway;  B ~= Server;
              Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB,
                              Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
-               : set evs4;
+               \\<in> set evs4;
              Gets B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
-               : set evs4 |]
-          ==> Says B A {|Nonce NA, X|} # evs4 : otway"
+               \\<in> set evs4 |]
+          ==> Says B A {|Nonce NA, X|} # evs4 \\<in> otway"
 
          (*This message models possible leaks of session keys.  The nonces
            identify the protocol run.*)
-    Oops "[| evso: otway;  
+    Oops "[| evso \\<in> otway;  
              Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
-               : set evso |]
-          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso : otway"
+               \\<in> set evso |]
+          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \\<in> otway"
 
 end
--- a/src/HOL/Auth/Recur.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Recur.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -133,27 +133,6 @@
 	Spy_analz_shrK RSN (2, rev_iffD1)];
 
 
-(** Nobody can have used non-existent keys! **)
-
-(*The special case of H={} has the same proof*)
-Goal "[| K \\<in> keysFor (parts (insert RB H));  RB \\<in> responses evs |] \
-\     ==> K \\<in> range shrK | K \\<in> keysFor (parts H)";
-by (etac rev_mp 1);
-by (etac responses.induct 1);
-by Auto_tac;
-qed_spec_mp "Key_in_keysFor_parts";
-
-
-Goal "evs \\<in> recur ==> Key K \\<notin> used evs --> K \\<notin> keysFor (parts (spies evs))";
-by (parts_induct_tac 1);
-(*RA3*)
-by (blast_tac (claset() addSDs [Key_in_keysFor_parts]) 2);
-(*Fake*)
-by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
-qed_spec_mp "new_keys_not_used";
-Addsimps [new_keys_not_used];
-
-
 
 (*** Proofs involving analz ***)
 
--- a/src/HOL/Auth/Recur.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Recur.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -19,21 +19,21 @@
 consts     respond :: "event list => (msg*msg*key)set"
 inductive "respond evs" (*Server's response to the nested message*)
   intrs
-    One  "[| Key KAB ~: used evs |]
+    One  "Key KAB \\<notin> used evs
           ==> (Hash[Key(shrK A)] {|Agent A, Agent B, Nonce NA, END|}, 
                {|Crypt (shrK A) {|Key KAB, Agent B, Nonce NA|}, END|},
-               KAB)   : respond evs"
+               KAB)   \\<in> respond evs"
 
     (*The most recent session key is passed up to the caller*)
-    Cons "[| (PA, RA, KAB) : respond evs;  
-             Key KBC ~: used evs;  Key KBC ~: parts {RA};
+    Cons "[| (PA, RA, KAB) \\<in> respond evs;  
+             Key KBC \\<notin> used evs;  Key KBC \\<notin> parts {RA};
              PA = Hash[Key(shrK A)] {|Agent A, Agent B, Nonce NA, P|} |]
           ==> (Hash[Key(shrK B)] {|Agent B, Agent C, Nonce NB, PA|}, 
                {|Crypt (shrK B) {|Key KBC, Agent C, Nonce NB|}, 
                  Crypt (shrK B) {|Key KAB, Agent A, Nonce NB|},
                  RA|},
                KBC)
-              : respond evs"
+              \\<in> respond evs"
 
 
 (*Induction over "respond" can be difficult due to the complexity of the
@@ -43,52 +43,52 @@
 inductive "responses evs"       
   intrs
     (*Server terminates lists*)
-    Nil  "END : responses evs"
+    Nil  "END \\<in> responses evs"
 
-    Cons "[| RA : responses evs;  Key KAB ~: used evs |]
+    Cons "[| RA \\<in> responses evs;  Key KAB \\<notin> used evs |]
           ==> {|Crypt (shrK B) {|Key KAB, Agent A, Nonce NB|},
-                RA|}  : responses evs"
+                RA|}  \\<in> responses evs"
 
 
 consts     recur   :: event list set
 inductive "recur"
   intrs 
          (*Initial trace is empty*)
-    Nil  "[]: recur"
+    Nil  "[] \\<in> recur"
 
          (*The spy MAY say anything he CAN say.  Common to
            all similar protocols.*)
     Fake "[| evs: recur;  X: synth (analz (spies evs)) |]
-          ==> Says Spy B X  # evs : recur"
+          ==> Says Spy B X  # evs \\<in> recur"
 
          (*Alice initiates a protocol run.
            END is a placeholder to terminate the nesting.*)
-    RA1  "[| evs1: recur;  Nonce NA ~: used evs1 |]
+    RA1  "[| evs1: recur;  Nonce NA \\<notin> used evs1 |]
           ==> Says A B (Hash[Key(shrK A)] {|Agent A, Agent B, Nonce NA, END|})
-              # evs1 : recur"
+              # evs1 \\<in> recur"
 
          (*Bob's response to Alice's message.  C might be the Server.
            We omit PA = {|XA, Agent A, Agent B, Nonce NA, P|} because
            it complicates proofs, so B may respond to any message at all!*)
-    RA2  "[| evs2: recur;  Nonce NB ~: used evs2;
-             Says A' B PA : set evs2 |]
+    RA2  "[| evs2: recur;  Nonce NB \\<notin> used evs2;
+             Says A' B PA \\<in> set evs2 |]
           ==> Says B C (Hash[Key(shrK B)] {|Agent B, Agent C, Nonce NB, PA|})
-              # evs2 : recur"
+              # evs2 \\<in> recur"
 
          (*The Server receives Bob's message and prepares a response.*)
-    RA3  "[| evs3: recur;  Says B' Server PB : set evs3;
-             (PB,RB,K) : respond evs3 |]
-          ==> Says Server B RB # evs3 : recur"
+    RA3  "[| evs3: recur;  Says B' Server PB \\<in> set evs3;
+             (PB,RB,K) \\<in> respond evs3 |]
+          ==> Says Server B RB # evs3 \\<in> recur"
 
          (*Bob receives the returned message and compares the Nonces with
            those in the message he previously sent the Server.*)
     RA4  "[| evs4: recur;  
              Says B  C {|XH, Agent B, Agent C, Nonce NB, 
-                         XA, Agent A, Agent B, Nonce NA, P|} : set evs4;
+                         XA, Agent A, Agent B, Nonce NA, P|} \\<in> set evs4;
              Says C' B {|Crypt (shrK B) {|Key KBC, Agent C, Nonce NB|}, 
                          Crypt (shrK B) {|Key KAB, Agent A, Nonce NB|}, 
-                         RA|} : set evs4 |]
-          ==> Says B A RA # evs4 : recur"
+                         RA|} \\<in> set evs4 |]
+          ==> Says B A RA # evs4 \\<in> recur"
 
 end
 
@@ -100,7 +100,7 @@
      the chain.  Oops cases proved using parts_cut, Key_in_keysFor_parts,
      etc.
 
-    Oops  "[| evso: recur;  Says Server B RB : set evso;
-	      RB : responses evs';  Key K : parts {RB} |]
-           ==> Notes Spy {|Key K, RB|} # evso : recur"
+    Oops  "[| evso: recur;  Says Server B RB \\<in> set evso;
+	      RB \\<in> responses evs';  Key K \\<in> parts {RB} |]
+           ==> Notes Spy {|Key K, RB|} # evso \\<in> recur"
   *)
--- a/src/HOL/Auth/Shared_lemmas.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Shared_lemmas.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -56,25 +56,12 @@
 qed "Spy_knows_Spy_bad";
 AddSIs [Spy_knows_Spy_bad];
 
-(*For not_bad_tac*)
+(*For case analysis on whether or not an agent is compromised*)
 Goal "[| Crypt (shrK A) X : analz (knows Spy evs);  A: bad |] \
-\              ==> X : analz (knows Spy evs)";
+\     ==> X : analz (knows Spy evs)";
 by (force_tac (claset() addSDs [analz.Decrypt], simpset()) 1);
 qed "Crypt_Spy_analz_bad";
 
-(*Prove that the agent is uncompromised by the confidentiality of 
-  a component of a message she's said.*)
-fun not_bad_tac s =
-    case_tac ("(" ^ s ^ ") : bad") THEN'
-    SELECT_GOAL 
-      (REPEAT_DETERM (etac exE 1) THEN
-       REPEAT_DETERM (dtac (Says_imp_spies RS analz.Inj) 1) THEN
-       REPEAT_DETERM (etac MPair_analz 1) THEN
-       THEN_BEST_FIRST 
-         (dres_inst_tac [("A", s)] Crypt_Spy_analz_bad 1 THEN assume_tac 1)
-         (has_fewer_prems 1, size_of_thm)
-         (Step_tac 1));
-
 
 (** Fresh keys never clash with long-term shared keys **)
 
--- a/src/HOL/Auth/TLS.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/TLS.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -32,13 +32,13 @@
 
 (*** clientK and serverK make symmetric keys; no clashes with pubK or priK ***)
 
-Goal "pubK A ~= sessionK arg";
+Goal "pubK A \\<noteq> sessionK arg";
 by (rtac notI 1);
 by (dres_inst_tac [("f","isSymKey")] arg_cong 1);
 by (Full_simp_tac 1);
 qed "pubK_neq_sessionK";
 
-Goal "priK A ~= sessionK arg";
+Goal "priK A \\<noteq> sessionK arg";
 by (rtac notI 1);
 by (dres_inst_tac [("f","isSymKey")] arg_cong 1);
 by (Full_simp_tac 1);
@@ -55,17 +55,17 @@
 
 
 (** These proofs assume that the Nonce_supply nonces 
-	(which have the form  @ N. Nonce N ~: used evs)
+	(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 "[| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF;  \
-\        A ~= B |]            \
-\     ==> EX SID M. EX evs: tls.    \
-\          Notes A {|Number SID, Agent A, Agent B, Nonce M|} : set evs";
+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
@@ -75,10 +75,10 @@
 result();
 
 (*And one for ServerAccepts.  Either FINISHED message may come first.*)
-Goal "[| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF;  \
-\        A ~= B |]                        \
-\     ==> EX SID NA PA NB PB M. EX evs: tls.    \
-\          Notes B {|Number SID, Agent A, Agent B, Nonce M|} : set evs";
+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
@@ -88,10 +88,10 @@
 result();
 
 (*Another one, for CertVerify (which is optional)*)
-Goal "[| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF;  \
-\        A ~= B |]                       \
-\  ==> EX NB PMS. EX evs: tls.   \
-\  Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|})) : set evs";
+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);
@@ -100,17 +100,17 @@
 result();
 
 (*Another one, for session resumption (both ServerResume and ClientResume) *)
-Goal "[| evs0 : tls;     \
-\        Notes A {|Number SID, Agent A, Agent B, Nonce M|} : set evs0; \
-\        Notes B {|Number SID, Agent A, Agent B, Nonce M|} : set evs0; \
-\        ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF;  \
-\        A ~= B |] \
-\     ==> EX NA PA NB PB X. EX evs: tls.    \
+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) : set evs  &  \
-\           Says B A (Crypt (serverK(NA,NB,M)) X) : set evs";
+\           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);
@@ -124,7 +124,7 @@
 
 
 (*Induction for regularity theorems.  If induction formula has the form
-   X ~: analz (spies evs) --> ... then it shortens the proof by discarding
+   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
@@ -133,17 +133,17 @@
     ALLGOALS Asm_simp_tac;
 
 
-(** Theorems of the form X ~: parts (spies evs) imply that NOBODY
+(** 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 : tls ==> (Key (priK A) : parts (spies evs)) = (A : bad)";
+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 : tls ==> (Key (priK A) : analz (spies evs)) = (A : bad)";
+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];
@@ -157,7 +157,7 @@
   little point in doing so: the loss of their private keys is a worse
   breach of security.*)
 Goalw [certificate_def]
-    "[| certificate B KB : parts (spies evs);  evs : tls |] ==> pubK B = KB";
+    "[| 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);
@@ -181,17 +181,17 @@
 
 (*** Properties of items found in Notes ***)
 
-Goal "[| Notes A {|Agent B, X|} : set evs;  evs : tls |]  \
-\     ==> Crypt (pubK B) X : parts (spies evs)";
+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))|} : set evs; \
-\        evs : tls |]    \
-\     ==> Crypt (pubK B) (Nonce PMS) : parts (spies evs)";
+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);
@@ -202,9 +202,9 @@
 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))|} : set evs;\
-\        evs : tls |]    \
-\     ==> Notes A {|Agent B, Nonce PMS|} : set evs";
+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*)
@@ -215,10 +215,10 @@
 (*** 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 : parts (spies evs);                          \
+Goal "[| X \\<in> parts (spies evs);                          \
 \        X = Crypt (priK A) (Hash{|nb, Agent B, pms|});  \
-\        evs : tls;  A ~: bad |]                         \
-\     ==> Says A B X : set evs";
+\        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);
@@ -226,20 +226,20 @@
 val lemma = result();
 
 (*Final version: B checks X using the distributed KA instead of priK A*)
-Goal "[| X : parts (spies evs);                            \
+Goal "[| X \\<in> parts (spies evs);                            \
 \        X = Crypt (invKey KA) (Hash{|nb, Agent B, pms|}); \
-\        certificate A KA : parts (spies evs);             \
-\        evs : tls;  A ~: bad |]                           \
-\     ==> Says A B X : set evs";
+\        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|}) \
-\          : parts (spies evs);                          \
-\        evs : tls;  A ~: bad |]                         \
-\     ==> Notes A {|Agent B, Nonce PMS|} : set evs";
+\          \\<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);
@@ -247,15 +247,15 @@
 
 (*Final version using the distributed KA instead of priK A*)
 Goal "[| Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|}) \
-\          : parts (spies evs);                             \
-\        certificate A KA : parts (spies evs);              \
-\        evs : tls;  A ~: bad |]                            \
-\     ==> Notes A {|Agent B, Nonce PMS|} : set evs";
+\          \\<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 : tls ==> Notes A {|Agent B, Nonce (PRF x)|} ~: set evs";
+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);
@@ -263,8 +263,8 @@
 Addsimps [no_Notes_A_PRF];
 
 
-Goal "[| Nonce (PRF (PMS,NA,NB)) : parts (spies evs);  evs : tls |]  \
-\     ==> Nonce PMS : parts (spies evs)";
+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*)
@@ -279,10 +279,10 @@
 (*** Unicity results for PMS, the pre-master-secret ***)
 
 (*PMS determines B.*)
-Goal "[| Crypt(pubK B)  (Nonce PMS) : parts (spies evs); \
-\        Crypt(pubK B') (Nonce PMS) : parts (spies evs); \
-\        Nonce PMS ~: analz (spies evs);                 \
-\        evs : tls |]                                          \
+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);
@@ -300,9 +300,9 @@
 **)
 
 (*In A's internal Note, PMS determines A and B.*)
-Goal "[| Notes A  {|Agent B,  Nonce PMS|} : set evs;  \
-\        Notes A' {|Agent B', Nonce PMS|} : set evs;  \
-\        evs : tls |]                               \
+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);
@@ -316,9 +316,9 @@
 
 (*Key compromise lemma needed to prove analz_image_keys.
   No collection of keys can help the spy get new private keys.*)
-Goal "evs : tls                                      \
-\     ==> ALL KK. (Key(priK B) : analz (Key`KK Un (spies evs))) = \
-\         (priK B : KK | B : bad)";
+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
@@ -329,25 +329,25 @@
 
 
 (*slightly speeds up the big simplification below*)
-Goal "KK <= range sessionK ==> priK B ~: KK";
+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 : analz (G Un H)) --> (X : analz H)  ==> \
-\     (X : analz (G Un H))  =  (X : analz H)";
+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:
-\ ALL Z. (Nonce N : analz (Key`(sessionK`Z) Un (spies evs))) = \
-\        (Nonce N : analz (spies evs))";
+\ \\<forall>Z. (Nonce N \\<in> analz (Key`(sessionK`Z) Un (spies evs))) = \
+\        (Nonce N \\<in> analz (spies evs))";
 **)
 
-Goal "evs : tls ==>                                    \
-\ ALL KK. KK <= range sessionK -->                     \
-\         (Nonce N : analz (Key`KK Un (spies evs))) = \
-\         (Nonce N : 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]));
@@ -363,9 +363,9 @@
 qed_spec_mp "analz_image_keys";
 
 (*Knowing some session keys is no help in getting new nonces*)
-Goal "evs : tls ==>          \
-\     Nonce N : analz (insert (Key (sessionK z)) (spies evs)) =  \
-\     (Nonce N : analz (spies evs))";
+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];
@@ -380,10 +380,10 @@
   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 ~: parts (spies evs);  \
+Goal "[| Nonce PMS \\<notin> parts (spies evs);  \
 \        K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role);  \
-\        evs : tls |]             \
-\  ==> Key K ~: parts (spies evs) & (ALL Y. Crypt K Y ~: parts (spies evs))";
+\        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);
@@ -399,15 +399,15 @@
 		simpset()) 1));
 val lemma = result();
 
-Goal "[| Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), role)) : parts (spies evs); \
-\        evs : tls |]             \
-\     ==> Nonce PMS : parts (spies evs)";
+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  \
-\          : parts (spies evs);  evs : tls |]             \
-\     ==> Nonce PMS : parts (spies evs)";
+\          \\<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";
 
@@ -416,9 +416,9 @@
   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 "[| ALL A. Says A Spy (Key (sessionK((NA,NB,M),role))) ~: set evs; \
-\        Nonce M ~: analz (spies evs);  evs : tls |]   \
-\     ==> Key (sessionK((NA,NB,M),role)) ~: parts (spies evs)";
+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*)
@@ -430,11 +430,11 @@
 
 
 (*If A sends ClientKeyExch to an honest B, then the PMS will stay secret.*)
-Goal "[| evs : tls;  A ~: bad;  B ~: bad |]           \
-\     ==> Notes A {|Agent B, Nonce PMS|} : set evs  -->   \
-\         Nonce PMS ~: analz (spies evs)";
+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 ~: range PRF*)
+(*ClientAccepts and ServerAccepts: because PMS \\<notin> range PRF*)
 by (REPEAT (Force_tac 6));
 (*ClientHello, ServerHello, ClientKeyExch, ServerResume: 
   mostly freshness reasoning*)
@@ -450,9 +450,9 @@
 
 (*If A sends ClientKeyExch to an honest B, then the MASTER SECRET
   will stay secret.*)
-Goal "[| evs : tls;  A ~: bad;  B ~: bad |]           \
-\     ==> Notes A {|Agent B, Nonce PMS|} : set evs  -->   \
-\         Nonce (PRF(PMS,NA,NB)) ~: analz (spies evs)";
+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, 
@@ -476,9 +476,9 @@
 
 (*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) : set evs;  \
-\        Notes A {|Agent B, Nonce PMS|} : set evs;   \
-\        evs : tls;  A' ~= Spy |]                \
+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);
@@ -496,11 +496,11 @@
 
 (*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|} : set evs;  \
-\        Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs;  \
-\        A ~: bad;  B ~: bad; \
-\        evs : tls |]   \
-\     ==> Key (clientK(Na,Nb,PRF(PMS,NA,NB))) ~: parts (spies evs)";
+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*)
@@ -519,9 +519,9 @@
 
 (*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) : set evs;  \
-\        Notes A {|Agent B, Nonce PMS|} : set evs;  \
-\        evs : tls;  A ~: bad;  B ~: bad;  B' ~= Spy |]                \
+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);
@@ -540,10 +540,10 @@
 
 (*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|} : set evs;                   \
-\        Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs; \
-\        A ~: bad;  B ~: bad;  evs : tls |]                          \
-\     ==> Key (serverK(Na,Nb,PRF(PMS,NA,NB))) ~: parts (spies evs)";
+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);
@@ -569,10 +569,10 @@
 \                     Nonce Na, Number PA, Agent A,    \
 \                     Nonce Nb, Number PB, Agent B|}); \
 \        M = PRF(PMS,NA,NB);                           \
-\        evs : tls;  A ~: bad;  B ~: bad |]            \
-\     ==> Says B Spy (Key(serverK(Na,Nb,M))) ~: set evs --> \
-\         Notes A {|Agent B, Nonce PMS|} : set evs --> \
-\         X : parts (spies evs) --> Says B A X : set evs";
+\        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);
@@ -587,11 +587,11 @@
   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 : tls;  A ~: bad;  B ~: bad |]     \
-\     ==> Says B Spy (Key(serverK(Na,Nb,M))) ~: set evs --> \
-\         Notes A {|Agent B, Nonce PMS|} : set evs -->              \
-\         Crypt (serverK(Na,Nb,M)) Y : parts (spies evs)  -->  \
-\         (EX A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) : set evs)";
+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])));
@@ -615,11 +615,11 @@
      ClientFinished, then B can then check the quoted values PA, PB, etc.
 ***)
 
-Goal "[| M = PRF(PMS,NA,NB);  evs : tls;  A ~: bad;  B ~: bad |] \
-\     ==> Says A Spy (Key(clientK(Na,Nb,M))) ~: set evs --> \
-\         Notes A {|Agent B, Nonce PMS|} : set evs -->               \
-\         Crypt (clientK(Na,Nb,M)) Y : parts (spies evs) -->         \
-\         Says A B (Crypt (clientK(Na,Nb,M)) Y) : set evs";
+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);
@@ -640,13 +640,13 @@
      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))) ~: set evs;\
-\        Says A' B (Crypt (clientK(Na,Nb,M)) Y) : set evs; \
-\        certificate A KA : parts (spies evs);       \
+\        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|}))\
-\          : set evs;                                                  \
-\        evs : tls;  A ~: bad;  B ~: bad |]                             \
-\     ==> Says A B (Crypt (clientK(Na,Nb,M)) Y) : set evs";
+\          \\<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";
--- a/src/HOL/Auth/TLS.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/TLS.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -82,14 +82,14 @@
          "[]: tls"
 
     Fake (*The spy, an active attacker, MAY say anything he CAN say.*)
-         "[| evs: tls;  X: synth (analz (spies evs)) |]
-          ==> Says Spy B X # evs : tls"
+         "[| 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: tls;
+         "[| 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 : tls"
+			   Key (sessionK((NA,NB,M),role)) |} # evsSK \\<in> tls"
 
     ClientHello
 	 (*(7.4.1.2)
@@ -97,40 +97,40 @@
 	   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 ~: 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: tls;  Nonce NA ~: used evsCH;  NA ~: 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  :  tls"
+	        # evsCH  \\<in>  tls"
 
     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: tls;  Nonce NB ~: used evsSH;  NB ~: 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|}
-	       : set evsSH |]
-          ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH  :  tls"
+	       \\<in> set evsSH |]
+          ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH  \\<in>  tls"
 
     Certificate
          (*SERVER (7.4.2) or CLIENT (7.4.6) CERTIFICATE.*)
-         "evsC: tls ==> Says B A (certificate B (pubK B)) # evsC  :  tls"
+         "evsC \\<in> tls ==> Says B A (certificate B (pubK B)) # evsC  \\<in>  tls"
 
     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 ~: 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: tls;  Nonce PMS ~: used evsCX;  PMS ~: range PRF;
-             Says B' A (certificate B KB) : 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  :  tls"
+	      # evsCX  \\<in>  tls"
 
     CertVerify
 	(*The optional Certificate Verify (7.4.8) message contains the
@@ -138,11 +138,11 @@
           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: tls;  
-             Says B' A {|Nonce NB, Number SID, Number PB|} : set evsCV;
-	     Notes A {|Agent B, Nonce PMS|} : 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  :  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).
@@ -153,101 +153,101 @@
           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~=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: tls;  
+         "[| evsCF \\<in> tls;  
 	     Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}
-	       : set evsCF;
-             Says B' A {|Nonce NB, Number SID, Number PB|} : set evsCF;
-             Notes A {|Agent B, Nonce PMS|} : 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 NB, Number PB, Agent B|}))
-              # evsCF  :  tls"
+              # evsCF  \\<in>  tls"
 
     ServerFinished
 	(*Keeping A' and A'' distinct means B cannot even check that the
           two messages originate from the same source. *)
-         "[| evsSF: tls;
+         "[| evsSF \\<in> tls;
 	     Says A' B  {|Agent A, Nonce NA, Number SID, Number PA|}
-	       : set evsSF;
-	     Says B  A  {|Nonce NB, Number SID, Number PB|} : set evsSF;
-	     Says A'' B (Crypt (pubK B) (Nonce PMS)) : 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 NB, Number PB, Agent B|}))
-              # evsSF  :  tls"
+              # evsSF  \\<in>  tls"
 
     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: tls;
-	     Notes A {|Agent B, Nonce PMS|} : set evsCA;
+         "[| 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 NB, Number PB, Agent B|};
-             Says A  B (Crypt (clientK(NA,NB,M)) X) : set evsCA;
-             Says B' A (Crypt (serverK(NA,NB,M)) X) : set evsCA |]
+             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  :  tls"
+             Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA  \\<in>  tls"
 
     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: tls;
-	     A ~= B;
-             Says A'' B (Crypt (pubK B) (Nonce PMS)) : set evsSA;
+         "[| 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 NB, Number PB, Agent B|};
-             Says B  A (Crypt (serverK(NA,NB,M)) X) : set evsSA;
-             Says A' B (Crypt (clientK(NA,NB,M)) X) : set evsSA |]
+             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  :  tls"
+             Notes B {|Number SID, Agent A, Agent B, Nonce M|} # evsSA  \\<in>  tls"
 
     ClientResume
          (*If A recalls the SESSION_ID, then she sends a FINISHED message
            using the new nonces and stored MASTER SECRET.*)
-         "[| evsCR: 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|} : set evsCR;
-             Notes A {|Number SID, Agent A, Agent B, Nonce M|} : 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 NB, Number PB, Agent B|}))
-              # evsCR  :  tls"
+              # evsCR  \\<in>  tls"
 
     ServerResume
          (*Resumption (7.3):  If B finds the SESSION_ID then he can send
            a FINISHED message using the recovered MASTER SECRET*)
-         "[| evsSR: 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|} : set evsSR;  
-             Notes B {|Number SID, Agent A, Agent B, Nonce M|} : 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 NB, Number PB, Agent B|})) # evsSR
-	        :  tls"
+	        \\<in>  tls"
 
     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 ~= 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: tls;  A ~= Spy;
-	     Says A B (Crypt (sessionK((NA,NB,M),role)) X) : set evso |]
-          ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso  :  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"
 
 end
--- a/src/HOL/Auth/WooLam.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/WooLam.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -15,8 +15,8 @@
 
 
 (*A "possibility property": there are traces that reach the end*)
-Goal "EX NB. EX evs: woolam.  \
-\           Says Server B (Crypt (shrK B) {|Agent A, Nonce NB|}) : set evs";
+Goal "\\<exists>NB. \\<exists>evs \\<in> woolam.  \
+\           Says Server B (Crypt (shrK B) {|Agent A, Nonce NB|}) \\<in> set evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (woolam.Nil RS woolam.WL1 RS woolam.WL2 RS woolam.WL3 RS 
           woolam.WL4 RS woolam.WL5) 2);
@@ -28,31 +28,31 @@
 
 (** For reasoning about the encrypted portion of messages **)
 
-Goal "Says A' B X : set evs ==> X : analz (spies evs)";
+Goal "Says A' B X \\<in> set evs ==> X \\<in> analz (spies evs)";
 by (etac (Says_imp_spies RS analz.Inj) 1);
 qed "WL4_analz_spies";
 
 bind_thm ("WL4_parts_spies",
           WL4_analz_spies RS (impOfSubs analz_subset_parts));
 
-(*For proving the easier theorems about X ~: parts (spies evs) *)
+(*For proving the easier theorems about X \\<notin> parts (spies evs) *)
 fun parts_induct_tac i = 
     etac woolam.induct i  THEN 
     ftac WL4_parts_spies (i+5)  THEN
     prove_simple_subgoals_tac 1;
 
 
-(** Theorems of the form X ~: parts (spies evs) imply that NOBODY
+(** Theorems of the form X \\<notin> parts (spies evs) imply that NOBODY
     sends messages containing X! **)
 
 (*Spy never sees another agent's shared key! (unless it's bad at start)*)
-Goal "evs : woolam ==> (Key (shrK A) : parts (spies evs)) = (A : bad)";
+Goal "evs \\<in> woolam ==> (Key (shrK A) \\<in> parts (spies evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (Blast_tac 1);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
-Goal "evs : woolam ==> (Key (shrK A) : analz (spies evs)) = (A : bad)";
+Goal "evs \\<in> woolam ==> (Key (shrK A) \\<in> analz (spies evs)) = (A \\<in> bad)";
 by Auto_tac;
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
@@ -67,9 +67,9 @@
 (*** WL4 ***)
 
 (*If the encrypted message appears then it originated with Alice*)
-Goal "[| Crypt (shrK A) (Nonce NB) : parts (spies evs);  \
-\        A ~: bad;  evs : woolam |]                      \
-\     ==> EX B. Says A B (Crypt (shrK A) (Nonce NB)) : set evs";
+Goal "[| Crypt (shrK A) (Nonce NB) \\<in> parts (spies evs);  \
+\        A \\<notin> bad;  evs \\<in> woolam |]                      \
+\     ==> \\<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
@@ -79,9 +79,9 @@
   Alice, then she originated that certificate.  But we DO NOT know that B
   ever saw it: the Spy may have rerouted the message to the Server.*)
 Goal "[| Says B' Server {|Agent A, Agent B, Crypt (shrK A) (Nonce NB)|} \
-\          : set evs;                                                   \
-\        A ~: bad;  evs : woolam |]                                     \
-\     ==> EX B. Says A B (Crypt (shrK A) (Nonce NB)) : set evs";
+\          \\<in> set evs;                                                   \
+\        A \\<notin> bad;  evs \\<in> woolam |]                                     \
+\     ==> \\<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \\<in> set evs";
 by (blast_tac (claset() addSIs [NB_Crypt_imp_Alice_msg]) 1);
 qed "Server_trusts_WL4";
 
@@ -91,10 +91,10 @@
 (*** WL5 ***)
 
 (*Server sent WL5 only if it received the right sort of message*)
-Goal "[| Says Server B (Crypt (shrK B) {|Agent A, NB|}) : set evs;      \
-\        evs : woolam |]                                                \
-\     ==> EX B'. Says B' Server {|Agent A, Agent B, Crypt (shrK A) NB|} \
-\            : set evs";
+Goal "[| Says Server B (Crypt (shrK B) {|Agent A, NB|}) \\<in> set evs;      \
+\        evs \\<in> woolam |]                                                \
+\     ==> \\<exists>B'. Says B' Server {|Agent A, Agent B, Crypt (shrK A) NB|} \
+\            \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
@@ -103,9 +103,9 @@
 AddDs [Server_sent_WL5];
 
 (*If the encrypted message appears then it originated with the Server!*)
-Goal "[| Crypt (shrK B) {|Agent A, NB|} : parts (spies evs);  \
-\        B ~: bad;  evs : woolam |]                           \
-\     ==> Says Server B (Crypt (shrK B) {|Agent A, NB|}) : set evs";
+Goal "[| Crypt (shrK B) {|Agent A, NB|} \\<in> parts (spies evs);  \
+\        B \\<notin> bad;  evs \\<in> woolam |]                           \
+\     ==> Says Server B (Crypt (shrK B) {|Agent A, NB|}) \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Blast_tac 1);
@@ -116,15 +116,15 @@
   But A may have sent the nonce to some other agent and it could have reached
   the Server via the Spy.*)
 Goal "[| Says S B (Crypt (shrK B) {|Agent A, Nonce NB|}): set evs; \
-\        A ~: bad;  B ~: bad;  evs : woolam  |]                  \
-\     ==> EX B. Says A B (Crypt (shrK A) (Nonce NB)) : set evs";
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> woolam  |]                  \
+\     ==> \\<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \\<in> set evs";
 by (blast_tac (claset() addSDs [NB_Crypt_imp_Server_msg]) 1);
 qed "B_trusts_WL5";
 
 
 (*B only issues challenges in response to WL1.  Not used.*)
-Goal "[| Says B A (Nonce NB) : set evs;  B ~= Spy;  evs : woolam |]  \
-\     ==> EX A'. Says A' B (Agent A) : set evs";
+Goal "[| Says B A (Nonce NB) \\<in> set evs;  B \\<noteq> Spy;  evs \\<in> woolam |]  \
+\     ==> \\<exists>A'. Says A' B (Agent A) \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
@@ -132,10 +132,10 @@
 
 
 (**CANNOT be proved because A doesn't know where challenges come from...
-Goal "[| A ~: bad;  B ~= Spy;  evs : woolam |]           \
-\ ==> Crypt (shrK A) (Nonce NB) : parts (spies evs) &  \
-\     Says B A (Nonce NB) : set evs                       \
-\     --> Says A B (Crypt (shrK A) (Nonce NB)) : set evs";
+Goal "[| A \\<notin> bad;  B \\<noteq> Spy;  evs \\<in> woolam |]           \
+\ ==> Crypt (shrK A) (Nonce NB) \\<in> parts (spies evs) &  \
+\     Says B A (Nonce NB) \\<in> set evs                       \
+\     --> Says A B (Crypt (shrK A) (Nonce NB)) \\<in> set evs";
 by (parts_induct_tac 1);
 by (Blast_tac 1);
 by Safe_tac;
--- a/src/HOL/Auth/WooLam.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/WooLam.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -20,46 +20,46 @@
 inductive woolam
   intrs 
          (*Initial trace is empty*)
-    Nil  "[]: woolam"
+    Nil  "[] \\<in> woolam"
 
          (** These rules allow agents to send messages to themselves **)
 
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evs: woolam;  X: synth (analz (spies evs)) |]
-          ==> Says Spy B X  # evs : woolam"
+    Fake "[| evsf \\<in> woolam;  X \\<in> synth (analz (spies evsf)) |]
+          ==> Says Spy B X  # evsf \\<in> woolam"
 
          (*Alice initiates a protocol run*)
-    WL1  "[| evs1: woolam |]
-          ==> Says A B (Agent A) # evs1 : woolam"
+    WL1  "[| evs1 \\<in> woolam |]
+          ==> Says A B (Agent A) # evs1 \\<in> woolam"
 
          (*Bob responds to Alice's message with a challenge.*)
-    WL2  "[| evs2: woolam;  Says A' B (Agent A) : set evs2 |]
-          ==> Says B A (Nonce NB) # evs2 : woolam"
+    WL2  "[| evs2 \\<in> woolam;  Says A' B (Agent A) \\<in> set evs2 |]
+          ==> Says B A (Nonce NB) # evs2 \\<in> woolam"
 
          (*Alice responds to Bob's challenge by encrypting NB with her key.
            B is *not* properly determined -- Alice essentially broadcasts
            her reply.*)
-    WL3  "[| evs3: woolam;
-             Says A  B (Agent A)  : set evs3;
-             Says B' A (Nonce NB) : set evs3 |]
-          ==> Says A B (Crypt (shrK A) (Nonce NB)) # evs3 : woolam"
+    WL3  "[| evs3 \\<in> woolam;
+             Says A  B (Agent A)  \\<in> set evs3;
+             Says B' A (Nonce NB) \\<in> set evs3 |]
+          ==> Says A B (Crypt (shrK A) (Nonce NB)) # evs3 \\<in> woolam"
 
          (*Bob forwards Alice's response to the Server.  NOTE: usually
            the messages are shown in chronological order, for clarity.
            But here, exchanging the two events would cause the lemma
            WL4_analz_spies to pick up the wrong assumption!*)
-    WL4  "[| evs4: woolam;  
-             Says A'  B X         : set evs4;
-             Says A'' B (Agent A) : set evs4 |]
-          ==> Says B Server {|Agent A, Agent B, X|} # evs4 : woolam"
+    WL4  "[| evs4 \\<in> woolam;  
+             Says A'  B X         \\<in> set evs4;
+             Says A'' B (Agent A) \\<in> set evs4 |]
+          ==> Says B Server {|Agent A, Agent B, X|} # evs4 \\<in> woolam"
 
          (*Server decrypts Alice's response for Bob.*)
-    WL5  "[| evs5: woolam;  
+    WL5  "[| evs5 \\<in> woolam;  
              Says B' Server {|Agent A, Agent B, Crypt (shrK A) (Nonce NB)|}
-               : set evs5 |]
+               \\<in> set evs5 |]
           ==> Says Server B (Crypt (shrK B) {|Agent A, Nonce NB|})
-                 # evs5 : woolam"
+                 # evs5 \\<in> woolam"
 
 end
--- a/src/HOL/Auth/Yahalom.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Yahalom.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -14,9 +14,9 @@
 
 
 (*A "possibility property": there are traces that reach the end*)
-Goal "A ~= Server \
-\     ==> EX X NB K. EX evs: yahalom.          \
-\            Says A B {|X, Crypt K (Nonce NB)|} : set evs";
+Goal "A \\<noteq> Server \
+\     ==> \\<exists>X NB K. \\<exists>evs \\<in> yahalom.          \
+\            Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (yahalom.Nil RS 
           yahalom.YM1 RS yahalom.Reception RS
@@ -25,21 +25,18 @@
 by possibility_tac;
 result();
 
-Goal "[| Gets B X : set evs; evs : yahalom |] ==> EX A. Says A B X : set evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> yahalom |] ==> \\<exists>A. Says A B X \\<in> set evs";
 by (etac rev_mp 1);
 by (etac yahalom.induct 1);
 by Auto_tac;
 qed "Gets_imp_Says";
 
 (*Must be proved separately for each protocol*)
-Goal "[| Gets B X : set evs; evs : yahalom |]  ==> X : knows Spy evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> yahalom |]  ==> X \\<in> knows Spy evs";
 by (blast_tac (claset() addSDs [Gets_imp_Says, Says_imp_knows_Spy]) 1);
 qed"Gets_imp_knows_Spy";
 AddDs [Gets_imp_knows_Spy RS parts.Inj];
 
-fun g_not_bad_tac s = 
-  ftac Gets_imp_Says THEN' assume_tac THEN' not_bad_tac s;
-
 
 (**** Inductive proofs about yahalom ****)
 
@@ -47,8 +44,8 @@
 (** For reasoning about the encrypted portion of messages **)
 
 (*Lets us treat YM4 using a similar argument as for the Fake case.*)
-Goal "[| Gets A {|Crypt (shrK A) Y, X|} : set evs;  evs : yahalom |]  \
-\     ==> X : analz (knows Spy evs)";
+Goal "[| Gets A {|Crypt (shrK A) Y, X|} \\<in> set evs;  evs \\<in> yahalom |]  \
+\     ==> X \\<in> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [Gets_imp_knows_Spy RS analz.Inj]) 1);
 qed "YM4_analz_knows_Spy";
 
@@ -56,13 +53,13 @@
           YM4_analz_knows_Spy RS (impOfSubs analz_subset_parts));
 
 (*For Oops*)
-Goal "Says Server A {|Crypt (shrK A) {|B,K,NA,NB|}, X|} : set evs \
-\     ==> K : parts (knows Spy evs)";
+Goal "Says Server A {|Crypt (shrK A) {|B,K,NA,NB|}, X|} \\<in> set evs \
+\     ==> K \\<in> parts (knows Spy evs)";
 by (blast_tac (claset() addSDs [parts.Body, 
                   Says_imp_knows_Spy RS parts.Inj]) 1);
 qed "YM4_Key_parts_knows_Spy";
 
-(*For proving the easier theorems about X ~: parts (knows Spy evs).*)
+(*For proving the easier theorems about X \\<notin> parts (knows Spy evs).*)
 fun parts_knows_Spy_tac i = 
   EVERY
    [ftac YM4_Key_parts_knows_Spy (i+7),
@@ -70,7 +67,7 @@
     prove_simple_subgoals_tac i];
 
 (*Induction for regularity theorems.  If induction formula has the form
-   X ~: analz (knows Spy evs) --> ... then it shortens the proof by discarding
+   X \\<notin> analz (knows Spy evs) --> ... then it shortens the proof by discarding
    needless information about analz (insert X (knows Spy evs))  *)
 fun parts_induct_tac i = 
     etac yahalom.induct i
@@ -79,18 +76,18 @@
     THEN  parts_knows_Spy_tac i;
 
 
-(** Theorems of the form X ~: parts (knows Spy evs) imply that NOBODY
+(** Theorems of the form X \\<notin> parts (knows Spy evs) imply that NOBODY
     sends messages containing X! **)
 
 (*Spy never sees another agent's shared key! (unless it's bad at start)*)
-Goal "evs : yahalom ==> (Key (shrK A) : parts (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> yahalom ==> (Key (shrK A) \\<in> parts (knows Spy evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 by (ALLGOALS Blast_tac);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
-Goal "evs : yahalom ==> (Key (shrK A) : analz (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> yahalom ==> (Key (shrK A) \\<in> analz (knows Spy evs)) = (A \\<in> bad)";
 by (auto_tac(claset() addDs [impOfSubs analz_subset_parts], simpset()));
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
@@ -100,27 +97,28 @@
 
 
 (*Nobody can have used non-existent keys!  Needed to apply analz_insert_Key*)
-Goal "evs : yahalom ==>          \
-\      Key K ~: used evs --> K ~: keysFor (parts (knows Spy evs))";
+Goal "evs \\<in> yahalom ==>          \
+\      Key K \\<notin> used evs --> K \\<notin> keysFor (parts (knows Spy evs))";
 by (parts_induct_tac 1);
 (*Fake*)
 by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
 (*YM2-4: Because Key K is not fresh, etc.*)
 by (REPEAT (blast_tac (claset() addSEs knows_Spy_partsEs) 1));
 qed_spec_mp "new_keys_not_used";
+Addsimps [new_keys_not_used];
 
+(*Earlier, \\<forall>protocol proofs declared this theorem.  
+  But Yahalom and Kerberos IV are the only ones that need it!*)
 bind_thm ("new_keys_not_analzd",
           [analz_subset_parts RS keysFor_mono,
            new_keys_not_used] MRS contra_subsetD);
 
-Addsimps [new_keys_not_used, new_keys_not_analzd];
-
 
 (*Describes the form of K when the Server sends this message.  Useful for
   Oops as well as main secrecy property.*)
 Goal "[| Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} \
-\          : set evs;   evs : yahalom |]                                \
-\     ==> K ~: range shrK";
+\          \\<in> set evs;   evs \\<in> yahalom |]                                \
+\     ==> K \\<notin> range shrK";
 by (etac rev_mp 1);
 by (etac yahalom.induct 1);
 by (ALLGOALS Asm_simp_tac);
@@ -137,18 +135,18 @@
 (****
  The following is to prove theorems of the form
 
-  Key K : analz (insert (Key KAB) (knows Spy evs)) ==>
-  Key K : analz (knows Spy evs)
+  Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) ==>
+  Key K \\<in> analz (knows Spy evs)
 
  A more general formula must be proved inductively.
 ****)
 
 (** Session keys are not used to encrypt other session keys **)
 
-Goal "evs : yahalom ==>                              \
-\  ALL K KK. KK <= - (range shrK) -->                \
-\         (Key K : analz (Key`KK Un (knows Spy evs))) = \
-\         (K : KK | Key K : analz (knows Spy evs))";
+Goal "evs \\<in> yahalom ==>                              \
+\  \\<forall>K KK. KK <= - (range shrK) -->                \
+\         (Key K \\<in> analz (Key`KK Un (knows Spy evs))) = \
+\         (K \\<in> KK | Key K \\<in> analz (knows Spy evs))";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (REPEAT_FIRST (resolve_tac [allI, impI]));
@@ -159,9 +157,9 @@
 by (spy_analz_tac 1);
 qed_spec_mp "analz_image_freshK";
 
-Goal "[| evs : yahalom;  KAB ~: range shrK |]              \
-\      ==> Key K : analz (insert (Key KAB) (knows Spy evs)) =  \
-\          (K = KAB | Key K : analz (knows Spy evs))";
+Goal "[| evs \\<in> yahalom;  KAB \\<notin> range shrK |]              \
+\      ==> Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) =  \
+\          (K = KAB | Key K \\<in> analz (knows Spy evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);
 qed "analz_insert_freshK";
 
@@ -170,10 +168,10 @@
 
 
 Goal "[| Says Server A                                                 \
-\         {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} : set evs; \
+\         {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} \\<in> set evs; \
 \       Says Server A'                                                \
-\         {|Crypt (shrK A') {|Agent B', Key K, na', nb'|}, X'|} : set evs; \
-\       evs : yahalom |]                                    \
+\         {|Crypt (shrK A') {|Agent B', Key K, na', nb'|}, X'|} \\<in> set evs; \
+\       evs \\<in> yahalom |]                                    \
 \    ==> A=A' & B=B' & na=na' & nb=nb'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -188,13 +186,13 @@
 
 (** Crucial secrecy property: Spy does not see the keys sent in msg YM3 **)
 
-Goal "[| A ~: bad;  B ~: bad;  evs : yahalom |]                \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
 \     ==> Says Server A                                        \
 \           {|Crypt (shrK A) {|Agent B, Key K, na, nb|},       \
 \             Crypt (shrK B) {|Agent A, Key K|}|}              \
-\          : set evs -->                                       \
-\         Notes Spy {|na, nb, Key K|} ~: set evs -->           \
-\         Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs -->                                       \
+\         Notes Spy {|na, nb, Key K|} \\<notin> set evs -->           \
+\         Key K \\<notin> analz (knows Spy evs)";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (ALLGOALS
@@ -216,10 +214,10 @@
 Goal "[| Says Server A                                         \
 \           {|Crypt (shrK A) {|Agent B, Key K, na, nb|},       \
 \             Crypt (shrK B) {|Agent A, Key K|}|}              \
-\          : set evs;                                          \
-\        Notes Spy {|na, nb, Key K|} ~: set evs;               \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                \
-\     ==> Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs;                                          \
+\        Notes Spy {|na, nb, Key K|} \\<notin> set evs;               \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSEs [lemma]) 1);
 qed "Spy_not_see_encrypted_key";
 
@@ -227,22 +225,22 @@
 (** Security Guarantee for A upon receiving YM3 **)
 
 (*If the encrypted message appears then it originated with the Server*)
-Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} : parts (knows Spy evs); \
-\        A ~: bad;  evs : yahalom |]                          \
+Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \\<in> parts (knows Spy evs); \
+\        A \\<notin> bad;  evs \\<in> yahalom |]                          \
 \      ==> Says Server A                                            \
 \           {|Crypt (shrK A) {|Agent B, Key K, na, nb|},            \
 \             Crypt (shrK B) {|Agent A, Key K|}|}                   \
-\          : set evs";
+\          \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 qed "A_trusts_YM3";
 
 (*The obvious combination of A_trusts_YM3 with Spy_not_see_encrypted_key*)
-Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} : parts (knows Spy evs); \
-\        Notes Spy {|na, nb, Key K|} ~: set evs;               \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                \
-\     ==> Key K ~: analz (knows Spy evs)";
+Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \\<in> parts (knows Spy evs); \
+\        Notes Spy {|na, nb, Key K|} \\<notin> set evs;               \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [A_trusts_YM3, Spy_not_see_encrypted_key]) 1);
 qed "A_gets_good_key";
 
@@ -250,13 +248,13 @@
 
 (*B knows, by the first part of A's message, that the Server distributed 
   the key for A and B.  But this part says nothing about nonces.*)
-Goal "[| Crypt (shrK B) {|Agent A, Key K|} : parts (knows Spy evs);      \
-\        B ~: bad;  evs : yahalom |]                                 \
-\     ==> EX NA NB. Says Server A                                    \
+Goal "[| Crypt (shrK B) {|Agent A, Key K|} \\<in> parts (knows Spy evs);      \
+\        B \\<notin> bad;  evs \\<in> yahalom |]                                 \
+\     ==> \\<exists>NA NB. Says Server A                                    \
 \                     {|Crypt (shrK A) {|Agent B, Key K,             \
 \                                        Nonce NA, Nonce NB|},       \
 \                       Crypt (shrK B) {|Agent A, Key K|}|}          \
-\                    : set evs";
+\                    \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
@@ -266,27 +264,27 @@
 
 (*B knows, by the second part of A's message, that the Server distributed 
   the key quoting nonce NB.  This part says nothing about agent names. 
-  Secrecy of NB is crucial.  Note that  Nonce NB ~: analz(knows Spy evs)  must
+  Secrecy of NB is crucial.  Note that  Nonce NB \\<notin> analz(knows Spy evs)  must
   be the FIRST antecedent of the induction formula.*)
-Goal "evs : yahalom                                          \
-\     ==> Nonce NB ~: analz (knows Spy evs) -->                  \
-\         Crypt K (Nonce NB) : parts (knows Spy evs) -->         \
-\         (EX A B NA. Says Server A                          \
+Goal "evs \\<in> yahalom                                          \
+\     ==> Nonce NB \\<notin> analz (knows Spy evs) -->                  \
+\         Crypt K (Nonce NB) \\<in> parts (knows Spy evs) -->         \
+\         (\\<exists>A B NA. Says Server A                          \
 \                     {|Crypt (shrK A) {|Agent B, Key K,     \
 \                               Nonce NA, Nonce NB|},        \
 \                       Crypt (shrK B) {|Agent A, Key K|}|}  \
-\                    : set evs)";
+\                    \\<in> set evs)";
 by (parts_induct_tac 1);
 by (ALLGOALS Clarify_tac);
 (*YM3 & Fake*)
 by (Blast_tac 2);
 by (Fake_parts_insert_tac 1);
 (*YM4*)
-(*A is uncompromised because NB is secure*)
-by (g_not_bad_tac "A" 1);
-(*A's certificate guarantees the existence of the Server message*)
-by (blast_tac (claset() addDs [Says_imp_knows_Spy RS parts.Inj RS parts.Fst RS
-			       A_trusts_YM3]) 1);
+(*A is uncompromised because NB is secure;
+  A's certificate guarantees the existence of the Server message*)
+by (blast_tac (claset() addSDs [Gets_imp_Says, Crypt_Spy_analz_bad]
+			addDs  [Says_imp_spies, analz.Inj, 
+			        parts.Inj RS parts.Fst RS A_trusts_YM3]) 1);
 bind_thm ("B_trusts_YM4_newK", result() RS mp RSN (2, rev_mp));
 
 
@@ -297,14 +295,14 @@
 Goalw [KeyWithNonce_def]
  "Says Server A                                              \
 \         {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB|}, X|} \
-\       : set evs ==> KeyWithNonce K NB evs";
+\       \\<in> set evs ==> KeyWithNonce K NB evs";
 by (Blast_tac 1);
 qed "KeyWithNonceI";
 
 Goalw [KeyWithNonce_def]
    "KeyWithNonce K NB (Says S A X # evs) =                                    \
 \ (Server = S &                                                            \
-\  (EX B n X'. X = {|Crypt (shrK A) {|Agent B, Key K, n, Nonce NB|}, X'|}) \
+\  (\\<exists>B n X'. X = {|Crypt (shrK A) {|Agent B, Key K, n, Nonce NB|}, X'|}) \
 \ | KeyWithNonce K NB evs)";
 by (Simp_tac 1);
 by (Blast_tac 1);
@@ -326,7 +324,7 @@
 (*A fresh key cannot be associated with any nonce 
   (with respect to a given trace). *)
 Goalw [KeyWithNonce_def]
- "Key K ~: used evs ==> ~ KeyWithNonce K NB evs";
+ "Key K \\<notin> used evs ==> ~ KeyWithNonce K NB evs";
 by (blast_tac (claset() addSEs knows_Spy_partsEs) 1);
 qed "fresh_not_KeyWithNonce";
 
@@ -335,8 +333,8 @@
 Goalw [KeyWithNonce_def]
  "[| Says Server A                                                \
 \             {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB'|}, X|} \
-\          : set evs;                                                 \
-\        NB ~= NB';  evs : yahalom |]                                 \
+\          \\<in> set evs;                                                 \
+\        NB \\<noteq> NB';  evs \\<in> yahalom |]                                 \
 \     ==> ~ KeyWithNonce K NB evs";
 by (blast_tac (claset() addDs [unique_session_keys]) 1);
 qed "Says_Server_KeyWithNonce";
@@ -349,39 +347,43 @@
 
 (*As with analz_image_freshK, we take some pains to express the property
   as a logical equivalence so that the simplifier can apply it.*)
-Goal "P --> (X : analz (G Un H)) --> (X : analz H)  ==> \
-\     P --> (X : analz (G Un H)) = (X : analz H)";
+Goal "P --> (X \\<in> analz (G Un H)) --> (X \\<in> analz H)  ==> \
+\     P --> (X \\<in> analz (G Un H)) = (X \\<in> analz H)";
 by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1);
 val Nonce_secrecy_lemma = result();
 
-Goal "evs : yahalom ==>                                      \
-\     (ALL KK. KK <= - (range shrK) -->                      \
-\          (ALL K: KK. ~ KeyWithNonce K NB evs)   -->        \
-\          (Nonce NB : analz (Key`KK Un (knows Spy evs))) =     \
-\          (Nonce NB : analz (knows Spy evs)))";
+Goal "evs \\<in> yahalom ==>                                      \
+\     (\\<forall>KK. KK <= - (range shrK) -->                      \
+\          (\\<forall>K \\<in> KK. ~ KeyWithNonce K NB evs)   -->        \
+\          (Nonce NB \\<in> analz (Key`KK Un (knows Spy evs))) =     \
+\          (Nonce NB \\<in> analz (knows Spy evs)))";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (REPEAT_FIRST (resolve_tac [impI RS allI]));
 by (REPEAT_FIRST (rtac Nonce_secrecy_lemma));
-(*For Oops, simplification proves NBa~=NB.  By Says_Server_KeyWithNonce,
+(*For Oops, simplification proves NBa\\<noteq>NB.  By Says_Server_KeyWithNonce,
   we get (~ KeyWithNonce K NB evs); then simplification can apply the
   induction hypothesis with KK = {K}.*)
 by (ALLGOALS  (*4 seconds*)
     (asm_simp_tac 
      (analz_image_freshK_ss 
        addsimps split_ifs
-       addsimps [all_conj_distrib, analz_image_freshK,
+       addsimps [all_conj_distrib, ball_conj_distrib, analz_image_freshK,
 		 KeyWithNonce_Says, KeyWithNonce_Notes, KeyWithNonce_Gets,
 		 fresh_not_KeyWithNonce, Says_Server_not_range,
-		 imp_disj_not1,		     (*Moves NBa~=NB to the front*)
+		 imp_disj_not1,		     (*Moves NBa\\<noteq>NB to the front*)
 		 Says_Server_KeyWithNonce])));
 (*Fake*) 
 by (spy_analz_tac 1);
 (*YM4*)  (** LEVEL 6 **)
-by (g_not_bad_tac "A" 1);
-by (dtac (Gets_imp_knows_Spy RS parts.Inj RS parts.Fst RS A_trusts_YM3) 1
-    THEN REPEAT (assume_tac 1));
-by (blast_tac (claset() addIs [KeyWithNonceI]) 1);
+by (thin_tac "\\<forall>KK. ?P KK" 1);
+by (Clarify_tac 1);  
+(*If A:bad then NBa is known, therefore NBa \\<noteq> NB.  Previous two steps make
+  the next step faster.*)
+by (blast_tac (claset() addSDs [Gets_imp_Says, Says_imp_spies, 
+                                Crypt_Spy_analz_bad]
+           addDs [analz.Inj,
+                  parts.Inj RS parts.Fst RS A_trusts_YM3 RS KeyWithNonceI]) 1);
 qed_spec_mp "Nonce_secrecy";
 
 
@@ -390,10 +392,10 @@
   for the induction to carry through.*)
 Goal "[| Says Server A                                               \
 \         {|Crypt (shrK A) {|Agent B, Key KAB, na, Nonce NB'|}, X|}  \
-\        : set evs;                                                  \
-\        NB ~= NB';  KAB ~: range shrK;  evs : yahalom |]            \
-\     ==> (Nonce NB : analz (insert (Key KAB) (knows Spy evs))) =        \
-\         (Nonce NB : analz (knows Spy evs))";
+\        \\<in> set evs;                                                  \
+\        NB \\<noteq> NB';  KAB \\<notin> range shrK;  evs \\<in> yahalom |]            \
+\     ==> (Nonce NB \\<in> analz (insert (Key KAB) (knows Spy evs))) =        \
+\         (Nonce NB \\<in> analz (knows Spy evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps 
 		  [Nonce_secrecy, Says_Server_KeyWithNonce]) 1);
 qed "single_Nonce_secrecy";
@@ -401,9 +403,9 @@
 
 (*** The Nonce NB uniquely identifies B's message. ***)
 
-Goal "[| Crypt (shrK B) {|Agent A, Nonce NA, nb|} : parts (knows Spy evs);    \
-\        Crypt (shrK B') {|Agent A', Nonce NA', nb|} : parts (knows Spy evs); \
-\       evs : yahalom;  B ~: bad;  B' ~: bad |]  \
+Goal "[| Crypt (shrK B) {|Agent A, Nonce NA, nb|} \\<in> parts (knows Spy evs);    \
+\        Crypt (shrK B') {|Agent A', Nonce NA', nb|} \\<in> parts (knows Spy evs); \
+\       evs \\<in> yahalom;  B \\<notin> bad;  B' \\<notin> bad |]  \
 \     ==> NA' = NA & A' = A & B' = B";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -414,27 +416,26 @@
 qed "unique_NB";
 
 
-(*Variant useful for proving secrecy of NB: the Says... form allows 
-  not_bad_tac to remove the assumption B' ~: bad.*)
+(*Variant useful for proving secrecy of NB.  Because nb is assumed to be 
+  secret, we no longer must assume B, B' not bad.*)
 Goal "[| Says C S   {|X,  Crypt (shrK B) {|Agent A, Nonce NA, nb|}|}    \
-\          : set evs;          B ~: bad;                                \
+\          \\<in> set evs;                          \
 \        Gets S' {|X', Crypt (shrK B') {|Agent A', Nonce NA', nb|}|}    \
-\          : set evs;                                                   \
-\        nb ~: analz (knows Spy evs);  evs : yahalom |]                 \
+\          \\<in> set evs;                                                   \
+\        nb \\<notin> analz (knows Spy evs);  evs \\<in> yahalom |]                 \
 \     ==> NA' = NA & A' = A & B' = B";
-by (g_not_bad_tac "B'" 1);
-by (blast_tac (claset() addSDs [Says_imp_knows_Spy RS parts.Inj]
-                        addSEs [MPair_parts]
-                        addDs  [unique_NB]) 1);
+by (blast_tac (claset() addSDs [Gets_imp_Says, Crypt_Spy_analz_bad]
+			addDs  [Says_imp_spies, unique_NB, parts.Inj, 
+                                analz.Inj]) 1);
 qed "Says_unique_NB";
 
 
 (** A nonce value is never used both as NA and as NB **)
 
-Goal "evs : yahalom                     \
-\ ==> Nonce NB ~: analz (knows Spy evs) -->    \
-\  Crypt (shrK B') {|Agent A', Nonce NB, nb'|} : parts(knows Spy evs) --> \
-\  Crypt (shrK B)  {|Agent A, na, Nonce NB|} ~: parts(knows Spy evs)";
+Goal "evs \\<in> yahalom                     \
+\ ==> Nonce NB \\<notin> analz (knows Spy evs) -->    \
+\  Crypt (shrK B') {|Agent A', Nonce NB, nb'|} \\<in> parts(knows Spy evs) --> \
+\  Crypt (shrK B)  {|Agent A, na, Nonce NB|} \\<notin> parts(knows Spy evs)";
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 by (blast_tac (claset() addDs [Gets_imp_knows_Spy RS analz.Inj]
@@ -447,10 +448,10 @@
 
 (*The Server sends YM3 only in response to YM2.*)
 Goal "[| Says Server A                                                \
-\         {|Crypt (shrK A) {|Agent B, k, na, nb|}, X|} : set evs;     \
-\        evs : yahalom |]                                             \
+\         {|Crypt (shrK A) {|Agent B, k, na, nb|}, X|} \\<in> set evs;     \
+\        evs \\<in> yahalom |]                                             \
 \     ==> Gets Server {| Agent B, Crypt (shrK B) {|Agent A, na, nb|} |} \
-\            : set evs";
+\            \\<in> set evs";
 by (etac rev_mp 1);
 by (etac yahalom.induct 1);
 by Auto_tac;
@@ -458,18 +459,18 @@
 
 
 (*A vital theorem for B, that nonce NB remains secure from the Spy.*)
-Goal "[| A ~: bad;  B ~: bad;  evs : yahalom |]  \
-\ ==> Says B Server                                                    \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]  \
+\ ==> (\\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \\<notin> set evs) -->      \
+\  Says B Server                                                    \
 \       {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|} \
-\  : set evs -->                                                    \
-\  (ALL k. Notes Spy {|Nonce NA, Nonce NB, k|} ~: set evs) -->      \
-\  Nonce NB ~: analz (knows Spy evs)";
+\  \\<in> set evs -->                                                    \
+\  Nonce NB \\<notin> analz (knows Spy evs)";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (ALLGOALS
     (asm_simp_tac 
      (simpset() addsimps split_ifs @ pushes @
-	                 [analz_insert_eq, analz_insert_freshK])));
+               [new_keys_not_analzd, analz_insert_eq, analz_insert_freshK])));
 (*Prove YM3 by showing that no NB can also be an NA*)
 by (blast_tac (claset() addDs [Says_imp_knows_Spy RS parts.Inj]
 	                addSEs [no_nonce_YM1_YM2, MPair_parts]
@@ -488,67 +489,58 @@
 by (ALLGOALS (Clarify_tac THEN' 
 	      full_simp_tac (simpset() addsimps [all_conj_distrib])));
 (*YM4: key K is visible to Spy, contradicting session key secrecy theorem*) 
-by (g_not_bad_tac "Aa" 1);
-by (dtac (Gets_imp_knows_Spy RS parts.Inj RS parts.Fst RS A_trusts_YM3) 1
-    THEN assume_tac 1);
-by (ftac Says_Server_imp_YM2 3);
-by (REPEAT_FIRST (eresolve_tac [asm_rl, exE]));
-(*  use Says_unique_NB to identify message components: Aa=A, Ba=B*)
-by (blast_tac (claset() addDs [Says_unique_NB, 
+(*Case analysis on Aa:bad; PROOF FAILED problems;
+  use Says_unique_NB to identify message components: Aa=A, Ba=B*)  
+by (blast_tac (claset() addSDs [Says_unique_NB, 
+                                parts.Inj RS parts.Fst RS A_trusts_YM3]
+			addDs [Gets_imp_knows_Spy RS analz.Inj, Gets_imp_Says,
+                               Says_imp_spies, Says_Server_imp_YM2,
 			       Spy_not_see_encrypted_key]) 1);
-(** LEVEL 13 **)
+(** LEVEL 9 **)
 (*Oops case: if the nonce is betrayed now, show that the Oops event is 
   covered by the quantified Oops assumption.*)
-by (ftac Says_Server_imp_YM2 1 THEN assume_tac 1);
-by (expand_case_tac "NB = NBa" 1);
+by (ftac Says_Server_imp_YM2 1 THEN assume_tac 1);  
+by (case_tac "NB = NBa" 1);
 (*If NB=NBa then all other components of the Oops message agree*)
 by (blast_tac (claset() addDs [Says_unique_NB]) 1);
-(*case NB ~= NBa*)
+(*case NB \\<noteq> NBa*)
 by (asm_simp_tac (simpset() addsimps [single_Nonce_secrecy]) 1);
-by (Clarify_tac 1);
-by (blast_tac (claset() addSEs [MPair_parts, no_nonce_YM1_YM2]
-                                             (*to prove NB~=NAa*)
+by (blast_tac (claset() addSEs [no_nonce_YM1_YM2] (*to prove NB\\<noteq>NAa*)
 		        addDs  [Says_imp_knows_Spy RS parts.Inj]) 1);
 bind_thm ("Spy_not_see_NB", result() RSN(2,rev_mp) RSN(2,rev_mp));
 
 
 (*B's session key guarantee from YM4.  The two certificates contribute to a
   single conclusion about the Server's message.  Note that the "Notes Spy"
-  assumption must quantify over ALL POSSIBLE keys instead of our particular K.
+  assumption must quantify over \\<forall>POSSIBLE keys instead of our particular K.
   If this run is broken and the spy substitutes a certificate containing an
   old key, B has no means of telling.*)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},                  \
-\                    Crypt K (Nonce NB)|} : set evs;                     \
+\                    Crypt K (Nonce NB)|} \\<in> set evs;                     \
 \        Says B Server                                                   \
 \          {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}   \
-\          : set evs;                                                    \
-\        ALL k. Notes Spy {|Nonce NA, Nonce NB, k|} ~: set evs;          \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]       \
+\          \\<in> set evs;                                                    \
+\        \\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \\<notin> set evs;          \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]       \
 \      ==> Says Server A                                                 \
 \                  {|Crypt (shrK A) {|Agent B, Key K,                    \
 \                            Nonce NA, Nonce NB|},                       \
 \                    Crypt (shrK B) {|Agent A, Key K|}|}                 \
-\            : set evs";
-by (ftac Spy_not_see_NB 1 THEN REPEAT (assume_tac 1));
-by (etac (Gets_imp_knows_Spy RS parts.Inj RS MPair_parts) 1 THEN
-    assume_tac 1 THEN dtac B_trusts_YM4_shrK 1);
-by (dtac B_trusts_YM4_newK 3);
-by (REPEAT_FIRST (eresolve_tac [asm_rl, exE]));
-by (ftac Says_Server_imp_YM2 1 THEN assume_tac 1);
-by (dtac unique_session_keys 1 THEN REPEAT (assume_tac 1));
-by (blast_tac (claset() addDs [Says_unique_NB]) 1);
+\            \\<in> set evs";
+by (blast_tac (claset() addDs [Spy_not_see_NB, Says_unique_NB,
+                               Says_Server_imp_YM2, B_trusts_YM4_newK]) 1);
 qed "B_trusts_YM4";
 
 
 (*The obvious combination of B_trusts_YM4 with Spy_not_see_encrypted_key*)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},                  \
-\                    Crypt K (Nonce NB)|} : set evs;                     \
+\                    Crypt K (Nonce NB)|} \\<in> set evs;                     \
 \        Says B Server                                                   \
 \          {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}   \
-\          : set evs;                                                    \
-\        ALL k. Notes Spy {|Nonce NA, Nonce NB, k|} ~: set evs;          \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                \
-\     ==> Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs;                                                    \
+\        \\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \\<notin> set evs;          \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [B_trusts_YM4, Spy_not_see_encrypted_key]) 1);
 qed "B_gets_good_key";
 
@@ -556,37 +548,37 @@
 (*** Authenticating B to A ***)
 
 (*The encryption in message YM2 tells us it cannot be faked.*)
-Goal "evs : yahalom                                            \
-\  ==> Crypt (shrK B) {|Agent A, Nonce NA, nb|} : parts (knows Spy evs) --> \
-\   B ~: bad -->                                              \
+Goal "evs \\<in> yahalom                                            \
+\  ==> Crypt (shrK B) {|Agent A, Nonce NA, nb|} \\<in> parts (knows Spy evs) --> \
+\   B \\<notin> bad -->                                              \
 \   Says B Server {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, nb|}|}  \
-\      : set evs";
+\      \\<in> set evs";
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 bind_thm ("B_Said_YM2", result() RSN (2, rev_mp) RS mp);
 
 (*If the server sends YM3 then B sent YM2*)
-Goal "evs : yahalom                                                      \
+Goal "evs \\<in> yahalom                                                      \
 \  ==> Says Server A {|Crypt (shrK A) {|Agent B, Key K, Nonce NA, nb|}, X|} \
-\      : set evs -->                                                     \
-\   B ~: bad -->                                                        \
+\      \\<in> set evs -->                                                     \
+\   B \\<notin> bad -->                                                        \
 \   Says B Server {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, nb|}|}  \
-\              : set evs";
+\              \\<in> set evs";
 by (etac yahalom.induct 1);
 by (ALLGOALS Asm_simp_tac);
 (*YM4*)
 by (Blast_tac 2);
-(*YM3 [blast_tac is 50% slower] *)
-by (best_tac (claset() addSDs [B_Said_YM2, Says_imp_knows_Spy RS parts.Inj]
-		       addSEs [MPair_parts]) 1);
+(*YM3*)
+by (blast_tac (claset() addSDs [B_Said_YM2, 
+                                Says_imp_knows_Spy RS parts.Inj]) 1);
 val lemma = result() RSN (2, rev_mp) RS mp |> standard;
 
 (*If A receives YM3 then B has used nonce NA (and therefore is alive)*)
 Goal "[| Gets A {|Crypt (shrK A) {|Agent B, Key K, Nonce NA, nb|}, X|} \
-\          : set evs;                                                    \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                        \
+\          \\<in> set evs;                                                    \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                        \
 \==> Says B Server {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, nb|}|} \
-\      : set evs";
+\      \\<in> set evs";
 by (blast_tac (claset() addSDs [A_trusts_YM3, lemma]
 		        addEs knows_Spy_partsEs) 1);
 qed "YM3_auth_B_to_A";
@@ -597,12 +589,12 @@
 (*Assuming the session key is secure, if both certificates are present then
   A has said NB.  We can't be sure about the rest of A's message, but only
   NB matters for freshness.*)  
-Goal "evs : yahalom                                             \
-\     ==> Key K ~: analz (knows Spy evs) -->                     \
-\         Crypt K (Nonce NB) : parts (knows Spy evs) -->         \
-\         Crypt (shrK B) {|Agent A, Key K|} : parts (knows Spy evs) --> \
-\         B ~: bad -->                                         \
-\         (EX X. Says A B {|X, Crypt K (Nonce NB)|} : set evs)";
+Goal "evs \\<in> yahalom                                             \
+\     ==> Key K \\<notin> analz (knows Spy evs) -->                     \
+\         Crypt K (Nonce NB) \\<in> parts (knows Spy evs) -->         \
+\         Crypt (shrK B) {|Agent A, Key K|} \\<in> parts (knows Spy evs) --> \
+\         B \\<notin> bad -->                                         \
+\         (\\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs)";
 by (parts_induct_tac 1);
 (*Fake*)
 by (Fake_parts_insert_tac 1);
@@ -611,30 +603,24 @@
 (*YM4: was Crypt K (Nonce NB) the very last message?  If not, use ind. hyp.*)
 by (asm_simp_tac (simpset() addsimps [ex_disj_distrib]) 1);
 (*yes: apply unicity of session keys*)
-by (g_not_bad_tac "Aa" 1);
-by (blast_tac (claset() addSEs [MPair_parts]
-                        addSDs [A_trusts_YM3, B_trusts_YM4_shrK]
-		        addDs  [Says_imp_knows_Spy RS parts.Inj,
-				unique_session_keys]) 1);
+by (blast_tac (claset() addSDs [Gets_imp_Says, A_trusts_YM3, B_trusts_YM4_shrK,
+                                Crypt_Spy_analz_bad]
+		addDs  [Says_imp_knows_Spy RS parts.Inj, 
+                        Says_imp_spies RS analz.Inj, unique_session_keys]) 1);
 qed_spec_mp "A_Said_YM3_lemma";
 
 (*If B receives YM4 then A has used nonce NB (and therefore is alive).
   Moreover, A associates K with NB (thus is talking about the same run).
   Other premises guarantee secrecy of K.*)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},                  \
-\                 Crypt K (Nonce NB)|} : set evs;                     \
+\                 Crypt K (Nonce NB)|} \\<in> set evs;                     \
 \        Says B Server                                                   \
 \          {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}   \
-\          : set evs;                                                    \
-\        (ALL NA k. Notes Spy {|Nonce NA, Nonce NB, k|} ~: set evs);     \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]       \
-\     ==> EX X. Says A B {|X, Crypt K (Nonce NB)|} : set evs";
-by (ftac B_trusts_YM4 1);
-by (REPEAT_FIRST (eresolve_tac [asm_rl, spec]));
-by (etac (Gets_imp_knows_Spy RS parts.Inj RS MPair_parts) 1 THEN assume_tac 1);
-by (rtac A_Said_YM3_lemma 1);
-by (rtac Spy_not_see_encrypted_key 2);
-by (REPEAT_FIRST assume_tac);
-by (blast_tac (claset() addSEs [MPair_parts]
-	       	        addDs [Says_imp_knows_Spy RS parts.Inj]) 1);
+\          \\<in> set evs;                                                    \
+\        (\\<forall>NA k. Notes Spy {|Nonce NA, Nonce NB, k|} \\<notin> set evs);     \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]       \
+\     ==> \\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs";
+by (blast_tac (claset() addSIs [A_Said_YM3_lemma]
+                   addDs [Spy_not_see_encrypted_key, B_trusts_YM4,
+                          Gets_imp_Says, Says_imp_knows_Spy RS parts.Inj]) 1);
 qed_spec_mp "YM4_imp_A_Said_YM3";
--- a/src/HOL/Auth/Yahalom.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Yahalom.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -16,65 +16,65 @@
 inductive "yahalom"
   intrs 
          (*Initial trace is empty*)
-    Nil  "[]: yahalom"
+    Nil  "[] \\<in> yahalom"
 
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evs: yahalom;  X: synth (analz (knows Spy evs)) |]
-          ==> Says Spy B X  # evs : yahalom"
+    Fake "[| evsf \\<in> yahalom;  X \\<in> synth (analz (knows Spy evsf)) |]
+          ==> Says Spy B X  # evsf \\<in> yahalom"
 
          (*A message that has been sent can be received by the
            intended recipient.*)
-    Reception "[| evsr: yahalom;  Says A B X : set evsr |]
-               ==> Gets B X # evsr : yahalom"
+    Reception "[| evsr \\<in> yahalom;  Says A B X \\<in> set evsr |]
+               ==> Gets B X # evsr \\<in> yahalom"
 
          (*Alice initiates a protocol run*)
-    YM1  "[| evs1: yahalom;  Nonce NA ~: used evs1 |]
-          ==> Says A B {|Agent A, Nonce NA|} # evs1 : yahalom"
+    YM1  "[| evs1 \\<in> yahalom;  Nonce NA \\<notin> used evs1 |]
+          ==> Says A B {|Agent A, Nonce NA|} # evs1 \\<in> yahalom"
 
          (*Bob's response to Alice's message.*)
-    YM2  "[| evs2: yahalom;  Nonce NB ~: used evs2;
-             Gets B {|Agent A, Nonce NA|} : set evs2 |]
+    YM2  "[| evs2 \\<in> yahalom;  Nonce NB \\<notin> used evs2;
+             Gets B {|Agent A, Nonce NA|} \\<in> set evs2 |]
           ==> Says B Server 
                   {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
-                # evs2 : yahalom"
+                # evs2 \\<in> yahalom"
 
          (*The Server receives Bob's message.  He responds by sending a
             new session key to Alice, with a packet for forwarding to Bob.*)
-    YM3  "[| evs3: yahalom;  Key KAB ~: used evs3;
+    YM3  "[| evs3 \\<in> yahalom;  Key KAB \\<notin> used evs3;
              Gets Server 
                   {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
-               : set evs3 |]
+               \\<in> set evs3 |]
           ==> Says Server A
                    {|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|},
                      Crypt (shrK B) {|Agent A, Key KAB|}|}
-                # evs3 : yahalom"
+                # evs3 \\<in> yahalom"
 
          (*Alice receives the Server's (?) message, checks her Nonce, and
            uses the new session key to send Bob his Nonce.  The premise
-           A ~= Server is needed to prove Says_Server_not_range.*)
-    YM4  "[| evs4: yahalom;  A ~= Server;
+           A \\<noteq> Server is needed to prove Says_Server_not_range.*)
+    YM4  "[| evs4 \\<in> yahalom;  A \\<noteq> Server;
              Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
-                : set evs4;
-             Says A B {|Agent A, Nonce NA|} : set evs4 |]
-          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 : yahalom"
+                \\<in> set evs4;
+             Says A B {|Agent A, Nonce NA|} \\<in> set evs4 |]
+          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 \\<in> yahalom"
 
          (*This message models possible leaks of session keys.  The Nonces
            identify the protocol run.  Quoting Server here ensures they are
            correct.*)
-    Oops "[| evso: yahalom;  
+    Oops "[| evso \\<in> yahalom;  
              Says Server A {|Crypt (shrK A)
                                    {|Agent B, Key K, Nonce NA, Nonce NB|},
-                             X|}  : set evso |]
-          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso : yahalom"
+                             X|}  \\<in> set evso |]
+          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \\<in> yahalom"
 
 
 constdefs 
   KeyWithNonce :: [key, nat, event list] => bool
   "KeyWithNonce K NB evs ==
-     EX A B na X. 
+     \\<exists>A B na X. 
        Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB|}, X|} 
-         : set evs"
+         \\<in> set evs"
 
 end
--- a/src/HOL/Auth/Yahalom2.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Yahalom2.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -17,8 +17,8 @@
 
 
 (*A "possibility property": there are traces that reach the end*)
-Goal "EX X NB K. EX evs: yahalom.          \
-\            Says A B {|X, Crypt K (Nonce NB)|} : set evs";
+Goal "\\<exists>X NB K. \\<exists>evs \\<in> yahalom.          \
+\            Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (yahalom.Nil RS 
           yahalom.YM1 RS yahalom.Reception RS
@@ -27,14 +27,14 @@
 by possibility_tac;
 result();
 
-Goal "[| Gets B X : set evs; evs : yahalom |] ==> EX A. Says A B X : set evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> yahalom |] ==> \\<exists>A. Says A B X \\<in> set evs";
 by (etac rev_mp 1);
 by (etac yahalom.induct 1);
 by Auto_tac;
 qed "Gets_imp_Says";
 
 (*Must be proved separately for each protocol*)
-Goal "[| Gets B X : set evs; evs : yahalom |]  ==> X : knows Spy evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> yahalom |]  ==> X \\<in> knows Spy evs";
 by (blast_tac (claset() addSDs [Gets_imp_Says, Says_imp_knows_Spy]) 1);
 qed"Gets_imp_knows_Spy";
 AddDs [Gets_imp_knows_Spy RS parts.Inj];
@@ -45,8 +45,8 @@
 (** For reasoning about the encrypted portion of messages **)
 
 (*Lets us treat YM4 using a similar argument as for the Fake case.*)
-Goal "[| Gets A {|NB, Crypt (shrK A) Y, X|} : set evs;  evs : yahalom |]  \
-\     ==> X : analz (knows Spy evs)";
+Goal "[| Gets A {|NB, Crypt (shrK A) Y, X|} \\<in> set evs;  evs \\<in> yahalom |]  \
+\     ==> X \\<in> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [Gets_imp_knows_Spy RS analz.Inj]) 1);
 qed "YM4_analz_knows_Spy";
 
@@ -54,13 +54,13 @@
           YM4_analz_knows_Spy RS (impOfSubs analz_subset_parts));
 
 (*For Oops*)
-Goal "Says Server A {|NB, Crypt (shrK A) {|B,K,NA|}, X|} : set evs \
-\     ==> K : parts (knows Spy evs)";
+Goal "Says Server A {|NB, Crypt (shrK A) {|B,K,NA|}, X|} \\<in> set evs \
+\     ==> K \\<in> parts (knows Spy evs)";
 by (blast_tac (claset() addSDs [parts.Body, 
          Says_imp_knows_Spy RS parts.Inj]) 1);
 qed "YM4_Key_parts_knows_Spy";
 
-(*For proving the easier theorems about X ~: parts (knows Spy evs).*)
+(*For proving the easier theorems about X \\<notin> parts (knows Spy evs).*)
 fun parts_knows_Spy_tac i = 
   EVERY
    [ftac YM4_Key_parts_knows_Spy (i+7),
@@ -68,7 +68,7 @@
     prove_simple_subgoals_tac i];
 
 (*Induction for regularity theorems.  If induction formula has the form
-   X ~: analz (knows Spy evs) --> ... then it shortens the proof by discarding
+   X \\<notin> analz (knows Spy evs) --> ... then it shortens the proof by discarding
    needless information about analz (insert X (knows Spy evs))  *)
 fun parts_induct_tac i = 
     etac yahalom.induct i
@@ -77,17 +77,17 @@
     THEN  parts_knows_Spy_tac i;
 
 
-(** Theorems of the form X ~: parts (knows Spy evs) imply that NOBODY
+(** Theorems of the form X \\<notin> parts (knows Spy evs) imply that NOBODY
     sends messages containing X! **)
 
 (*Spy never sees another agent's shared key! (unless it's bad at start)*)
-Goal "evs : yahalom ==> (Key (shrK A) : parts (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> yahalom ==> (Key (shrK A) \\<in> parts (knows Spy evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
-Goal "evs : yahalom ==> (Key (shrK A) : analz (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> yahalom ==> (Key (shrK A) \\<in> analz (knows Spy evs)) = (A \\<in> bad)";
 by Auto_tac;
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
@@ -97,8 +97,8 @@
 
 
 (*Nobody can have used non-existent keys!  Needed to apply analz_insert_Key*)
-Goal "evs : yahalom ==>          \
-\      Key K ~: used evs --> K ~: keysFor (parts (knows Spy evs))";
+Goal "evs \\<in> yahalom ==>          \
+\      Key K \\<notin> used evs --> K \\<notin> keysFor (parts (knows Spy evs))";
 by (parts_induct_tac 1);
 (*YM4: Key K is not fresh!*)
 by (Blast_tac 3);
@@ -107,19 +107,20 @@
 (*Fake*)
 by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
 qed_spec_mp "new_keys_not_used";
+Addsimps [new_keys_not_used];
 
+(*Earlier, ALL protocol proofs declared this theorem.  
+  But Yahalom and Kerberos IV are the only ones that need it!*)
 bind_thm ("new_keys_not_analzd",
           [analz_subset_parts RS keysFor_mono,
            new_keys_not_used] MRS contra_subsetD);
 
-Addsimps [new_keys_not_used, new_keys_not_analzd];
-
 (*Describes the form of K when the Server sends this message.  Useful for
   Oops as well as main secrecy property.*)
 Goal "[| Says Server A {|nb', Crypt (shrK A) {|Agent B, Key K, na|}, X|} \
-\         : set evs;                                            \
-\        evs : yahalom |]                                       \
-\     ==> K ~: range shrK";
+\         \\<in> set evs;                                            \
+\        evs \\<in> yahalom |]                                       \
+\     ==> K \\<notin> range shrK";
 by (etac rev_mp 1);
 by (etac yahalom.induct 1);
 by (ALLGOALS Asm_simp_tac);
@@ -137,8 +138,8 @@
 (****
  The following is to prove theorems of the form
 
-          Key K : analz (insert (Key KAB) (knows Spy evs)) ==>
-          Key K : analz (knows Spy evs)
+          Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) ==>
+          Key K \\<in> analz (knows Spy evs)
 
  A more general formula must be proved inductively.
 
@@ -146,10 +147,10 @@
 
 (** Session keys are not used to encrypt other session keys **)
 
-Goal "evs : yahalom ==>                               \
-\  ALL K KK. KK <= - (range shrK) -->                 \
-\         (Key K : analz (Key`KK Un (knows Spy evs))) =  \
-\         (K : KK | Key K : analz (knows Spy evs))";
+Goal "evs \\<in> yahalom ==>                               \
+\  \\<forall>K KK. KK <= - (range shrK) -->                 \
+\         (Key K \\<in> analz (Key`KK Un (knows Spy evs))) =  \
+\         (K \\<in> KK | Key K \\<in> analz (knows Spy evs))";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (REPEAT_FIRST (resolve_tac [allI, impI]));
@@ -159,9 +160,9 @@
 by (spy_analz_tac 1);
 qed_spec_mp "analz_image_freshK";
 
-Goal "[| evs : yahalom;  KAB ~: range shrK |] ==>     \
-\     Key K : analz (insert (Key KAB) (knows Spy evs)) =  \
-\     (K = KAB | Key K : analz (knows Spy evs))";
+Goal "[| evs \\<in> yahalom;  KAB \\<notin> range shrK |] ==>     \
+\     Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) =  \
+\     (K = KAB | Key K \\<in> analz (knows Spy evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);
 qed "analz_insert_freshK";
 
@@ -169,10 +170,10 @@
 (*** The Key K uniquely identifies the Server's  message. **)
 
 Goal "[| Says Server A                                            \
-\         {|nb, Crypt (shrK A) {|Agent B, Key K, na|}, X|} : set evs; \
+\         {|nb, Crypt (shrK A) {|Agent B, Key K, na|}, X|} \\<in> set evs; \
 \       Says Server A'                                           \
-\         {|nb', Crypt (shrK A') {|Agent B', Key K, na'|}, X'|} : set evs; \
-\       evs : yahalom |]                                         \
+\         {|nb', Crypt (shrK A') {|Agent B', Key K, na'|}, X'|} \\<in> set evs; \
+\       evs \\<in> yahalom |]                                         \
 \    ==> A=A' & B=B' & na=na' & nb=nb'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -185,23 +186,23 @@
 
 (** Crucial secrecy property: Spy does not see the keys sent in msg YM3 **)
 
-Goal "[| A ~: bad;  B ~: bad;  evs : yahalom |]              \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]              \
 \     ==> Says Server A                                      \
 \           {|nb, Crypt (shrK A) {|Agent B, Key K, na|},     \
 \                 Crypt (shrK B) {|Agent A, Agent B, Key K, nb|}|} \
-\          : set evs -->                                     \
-\         Notes Spy {|na, nb, Key K|} ~: set evs -->         \
-\         Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs -->                                     \
+\         Notes Spy {|na, nb, Key K|} \\<notin> set evs -->         \
+\         Key K \\<notin> analz (knows Spy evs)";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (ALLGOALS
     (asm_simp_tac 
      (simpset() addsimps split_ifs
-	        addsimps [analz_insert_eq, analz_insert_freshK])));
+	        addsimps [new_keys_not_analzd, analz_insert_eq, 
+                          analz_insert_freshK])));
 (*Oops*)
 by (blast_tac (claset() addDs [unique_session_keys]) 3);
-(*YM3: delete a useless induction hypothesis*)
-by (thin_tac "?P-->?Q" 2);
+(*YM3*)
 by (Blast_tac 2);
 (*Fake*) 
 by (spy_analz_tac 1);
@@ -212,12 +213,11 @@
 Goal "[| Says Server A                                    \
 \           {|nb, Crypt (shrK A) {|Agent B, Key K, na|},  \
 \                 Crypt (shrK B) {|Agent A, Agent B, Key K, nb|}|}    \
-\        : set evs;                                       \
-\        Notes Spy {|na, nb, Key K|} ~: set evs;          \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]           \
-\     ==> Key K ~: analz (knows Spy evs)";
-by (ftac Says_Server_message_form 1 THEN assume_tac 1);
-by (blast_tac (claset() addSEs [lemma]) 1);
+\        \\<in> set evs;                                       \
+\        Notes Spy {|na, nb, Key K|} \\<notin> set evs;          \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]           \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
+by (blast_tac (claset() addSEs [lemma] addDs [Says_Server_message_form]) 1);
 qed "Spy_not_see_encrypted_key";
 
 
@@ -226,22 +226,22 @@
 (*If the encrypted message appears then it originated with the Server.
   May now apply Spy_not_see_encrypted_key, subject to its conditions.*)
 Goal "[| Crypt (shrK A) {|Agent B, Key K, na|}                      \
-\         : parts (knows Spy evs);                                      \
-\        A ~: bad;  evs : yahalom |]                                \
-\      ==> EX nb. Says Server A                                     \
+\         \\<in> parts (knows Spy evs);                                      \
+\        A \\<notin> bad;  evs \\<in> yahalom |]                                \
+\      ==> \\<exists>nb. Says Server A                                     \
 \                   {|nb, Crypt (shrK A) {|Agent B, Key K, na|},    \
 \                         Crypt (shrK B) {|Agent A, Agent B, Key K, nb|}|} \
-\                 : set evs";
+\                 \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
 qed "A_trusts_YM3";
 
 (*The obvious combination of A_trusts_YM3 with Spy_not_see_encrypted_key*)
-Goal "[| Crypt (shrK A) {|Agent B, Key K, na|} : parts (knows Spy evs); \
-\        ALL nb. Notes Spy {|na, nb, Key K|} ~: set evs;            \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                     \
-\     ==> Key K ~: analz (knows Spy evs)";
+Goal "[| Crypt (shrK A) {|Agent B, Key K, na|} \\<in> parts (knows Spy evs); \
+\        \\<forall>nb. Notes Spy {|na, nb, Key K|} \\<notin> set evs;            \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                     \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [A_trusts_YM3, Spy_not_see_encrypted_key]) 1);
 qed "A_gets_good_key";
 
@@ -251,13 +251,13 @@
 (*B knows, by the first part of A's message, that the Server distributed 
   the key for A and B, and has associated it with NB.*)
 Goal "[| Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|} \
-\          : parts (knows Spy evs);                               \
-\        B ~: bad;  evs : yahalom |]                          \
-\ ==> EX NA. Says Server A                                       \
+\          \\<in> parts (knows Spy evs);                               \
+\        B \\<notin> bad;  evs \\<in> yahalom |]                          \
+\ ==> \\<exists>NA. Says Server A                                       \
 \            {|Nonce NB,                                      \
 \              Crypt (shrK A) {|Agent B, Key K, Nonce NA|},   \
 \              Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}|} \
-\            : set evs";
+\            \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (ALLGOALS Blast_tac);
@@ -271,13 +271,13 @@
   because we do not have to show that NB is secret. *)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}, \
 \                    X|}                                         \
-\          : set evs;                                            \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                  \
-\ ==> EX NA. Says Server A                                          \
+\          \\<in> set evs;                                            \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                  \
+\ ==> \\<exists>NA. Says Server A                                          \
 \            {|Nonce NB,                                         \
 \              Crypt (shrK A) {|Agent B, Key K, Nonce NA|},      \
 \              Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}|} \
-\           : set evs";
+\           \\<in> set evs";
 by (blast_tac (claset() addSDs [B_trusts_YM4_shrK]) 1);
 qed "B_trusts_YM4";
 
@@ -285,10 +285,10 @@
 (*The obvious combination of B_trusts_YM4 with Spy_not_see_encrypted_key*)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}, \
 \                    X|}                                         \
-\          : set evs;                                            \
-\        ALL na. Notes Spy {|na, Nonce NB, Key K|} ~: set evs;   \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                  \
-\     ==> Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs;                                            \
+\        \\<forall>na. Notes Spy {|na, Nonce NB, Key K|} \\<notin> set evs;   \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                  \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [B_trusts_YM4, Spy_not_see_encrypted_key]) 1);
 qed "B_gets_good_key";
 
@@ -297,11 +297,11 @@
 (*** Authenticating B to A ***)
 
 (*The encryption in message YM2 tells us it cannot be faked.*)
-Goal "[| Crypt (shrK B) {|Agent A, Nonce NA|} : parts (knows Spy evs);  \
-\        B ~: bad;  evs : yahalom                                   \
-\     |] ==> EX NB. Says B Server {|Agent B, Nonce NB,              \
+Goal "[| Crypt (shrK B) {|Agent A, Nonce NA|} \\<in> parts (knows Spy evs);  \
+\        B \\<notin> bad;  evs \\<in> yahalom                                   \
+\     |] ==> \\<exists>NB. Says B Server {|Agent B, Nonce NB,              \
 \                            Crypt (shrK B) {|Agent A, Nonce NA|}|} \
-\                     : set evs";
+\                     \\<in> set evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
@@ -312,11 +312,11 @@
 (*If the server sends YM3 then B sent YM2, perhaps with a different NB*)
 Goal "[| Says Server A                                              \
 \            {|nb, Crypt (shrK A) {|Agent B, Key K, Nonce NA|}, X|} \
-\          : set evs;                                               \
-\        B ~: bad;  evs : yahalom                                   \
-\     |] ==> EX nb'. Says B Server {|Agent B, nb',                  \
+\          \\<in> set evs;                                               \
+\        B \\<notin> bad;  evs \\<in> yahalom                                   \
+\     |] ==> \\<exists>nb'. Says B Server {|Agent B, nb',                  \
 \                            Crypt (shrK B) {|Agent A, Nonce NA|}|} \
-\                      : set evs";
+\                      \\<in> set evs";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
 by (etac yahalom.induct 1);
@@ -329,11 +329,11 @@
 
 (*If A receives YM3 then B has used nonce NA (and therefore is alive)*)
 Goal "[| Gets A {|nb, Crypt (shrK A) {|Agent B, Key K, Nonce NA|}, X|}   \
-\          : set evs;                                                    \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                          \
-\==> EX nb'. Says B Server                                               \
+\          \\<in> set evs;                                                    \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                          \
+\==> \\<exists>nb'. Says B Server                                               \
 \                 {|Agent B, nb', Crypt (shrK B) {|Agent A, Nonce NA|}|} \
-\              : set evs";
+\              \\<in> set evs";
 by (blast_tac (claset() addSDs [A_trusts_YM3, lemma]) 1);
 qed "YM3_auth_B_to_A";
 
@@ -342,15 +342,15 @@
 
 (*Assuming the session key is secure, if both certificates are present then
   A has said NB.  We can't be sure about the rest of A's message, but only
-  NB matters for freshness.  Note that  Key K ~: analz (knows Spy evs)  must be
+  NB matters for freshness.  Note that  Key K \\<notin> analz (knows Spy evs)  must be
   the FIRST antecedent of the induction formula.*)  
-Goal "evs : yahalom                                     \
-\     ==> Key K ~: analz (knows Spy evs) -->                \
-\         Crypt K (Nonce NB) : parts (knows Spy evs) -->    \
+Goal "evs \\<in> yahalom                                     \
+\     ==> Key K \\<notin> analz (knows Spy evs) -->                \
+\         Crypt K (Nonce NB) \\<in> parts (knows Spy evs) -->    \
 \         Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}      \
-\           : parts (knows Spy evs) -->                     \
-\         B ~: bad -->                                  \
-\         (EX X. Says A B {|X, Crypt K (Nonce NB)|} : set evs)";
+\           \\<in> parts (knows Spy evs) -->                     \
+\         B \\<notin> bad -->                                  \
+\         (\\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs)";
 by (parts_induct_tac 1);
 (*Fake*)
 by (Blast_tac 1);
@@ -358,12 +358,10 @@
 by (force_tac (claset() addSDs [Crypt_imp_keysFor], simpset()) 1); 
 (*YM4: was Crypt K (Nonce NB) the very last message?  If not, use ind. hyp.*)
 by (asm_simp_tac (simpset() addsimps [ex_disj_distrib]) 1);
-(*yes: delete a useless induction hypothesis; apply unicity of session keys*)
-by (thin_tac "?P-->?Q" 1);
-by (dtac Gets_imp_Says 1 THEN assume_tac 1);
-by (not_bad_tac "Aa" 1);
-by (blast_tac (claset() addSDs [A_trusts_YM3, B_trusts_YM4_shrK]
-			addDs  [unique_session_keys]) 1);
+(*Yes: apply unicity of session keys.  [Ind. hyp. no longer needed!]*)
+by (blast_tac (claset() addSDs [Gets_imp_Says, A_trusts_YM3, B_trusts_YM4_shrK,
+                                Crypt_Spy_analz_bad]
+		addDs  [Says_imp_spies RS analz.Inj, unique_session_keys]) 1);
 qed_spec_mp "Auth_A_to_B_lemma";
 
 
@@ -371,12 +369,10 @@
   Moreover, A associates K with NB (thus is talking about the same run).
   Other premises guarantee secrecy of K.*)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Agent B, Key K, Nonce NB|}, \
-\                    Crypt K (Nonce NB)|} : set evs;                 \
-\        (ALL NA. Notes Spy {|Nonce NA, Nonce NB, Key K|} ~: set evs); \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                    \
-\     ==> EX X. Says A B {|X, Crypt K (Nonce NB)|} : set evs";
-by (subgoal_tac "Key K ~: analz (knows Spy evs)" 1);
-by (blast_tac (claset() addIs [Auth_A_to_B_lemma]) 1);
-by (blast_tac (claset() addDs  [Spy_not_see_encrypted_key,
-				B_trusts_YM4_shrK]) 1);
+\                    Crypt K (Nonce NB)|} \\<in> set evs;                 \
+\        (\\<forall>NA. Notes Spy {|Nonce NA, Nonce NB, Key K|} \\<notin> set evs); \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                    \
+\     ==> \\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs";
+by (blast_tac (claset() addIs [Auth_A_to_B_lemma]
+                     addDs [Spy_not_see_encrypted_key, B_trusts_YM4_shrK]) 1);
 qed_spec_mp "YM4_imp_A_Said_YM3";
--- a/src/HOL/Auth/Yahalom2.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Yahalom2.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -19,58 +19,58 @@
 inductive "yahalom"
   intrs 
          (*Initial trace is empty*)
-    Nil  "[]: yahalom"
+    Nil  "[] \\<in> yahalom"
 
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evs: yahalom;  X: synth (analz (knows Spy evs)) |]
-          ==> Says Spy B X  # evs : yahalom"
+    Fake "[| evsf \\<in> yahalom;  X \\<in> synth (analz (knows Spy evsf)) |]
+          ==> Says Spy B X  # evsf \\<in> yahalom"
 
          (*A message that has been sent can be received by the
            intended recipient.*)
-    Reception "[| evsr: yahalom;  Says A B X : set evsr |]
-               ==> Gets B X # evsr : yahalom"
+    Reception "[| evsr \\<in> yahalom;  Says A B X \\<in> set evsr |]
+               ==> Gets B X # evsr \\<in> yahalom"
 
          (*Alice initiates a protocol run*)
-    YM1  "[| evs1: yahalom;  Nonce NA ~: used evs1 |]
-          ==> Says A B {|Agent A, Nonce NA|} # evs1 : yahalom"
+    YM1  "[| evs1 \\<in> yahalom;  Nonce NA \\<notin> used evs1 |]
+          ==> Says A B {|Agent A, Nonce NA|} # evs1 \\<in> yahalom"
 
          (*Bob's response to Alice's message.*)
-    YM2  "[| evs2: yahalom;  Nonce NB ~: used evs2;
-             Gets B {|Agent A, Nonce NA|} : set evs2 |]
+    YM2  "[| evs2 \\<in> yahalom;  Nonce NB \\<notin> used evs2;
+             Gets B {|Agent A, Nonce NA|} \\<in> set evs2 |]
           ==> Says B Server 
                   {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
-                # evs2 : yahalom"
+                # evs2 \\<in> yahalom"
 
          (*The Server receives Bob's message.  He responds by sending a
            new session key to Alice, with a certificate for forwarding to Bob.
            Both agents are quoted in the 2nd certificate to prevent attacks!*)
-    YM3  "[| evs3: yahalom;  Key KAB ~: used evs3;
+    YM3  "[| evs3 \\<in> yahalom;  Key KAB \\<notin> used evs3;
              Gets Server {|Agent B, Nonce NB,
 			   Crypt (shrK B) {|Agent A, Nonce NA|}|}
-               : set evs3 |]
+               \\<in> set evs3 |]
           ==> Says Server A
                {|Nonce NB, 
                  Crypt (shrK A) {|Agent B, Key KAB, Nonce NA|},
                  Crypt (shrK B) {|Agent A, Agent B, Key KAB, Nonce NB|}|}
-                 # evs3 : yahalom"
+                 # evs3 \\<in> yahalom"
 
          (*Alice receives the Server's (?) message, checks her Nonce, and
            uses the new session key to send Bob his Nonce.*)
-    YM4  "[| evs4: yahalom;  
+    YM4  "[| evs4 \\<in> yahalom;  
              Gets A {|Nonce NB, Crypt (shrK A) {|Agent B, Key K, Nonce NA|},
-                      X|}  : set evs4;
-             Says A B {|Agent A, Nonce NA|} : set evs4 |]
-          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 : yahalom"
+                      X|}  \\<in> set evs4;
+             Says A B {|Agent A, Nonce NA|} \\<in> set evs4 |]
+          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 \\<in> yahalom"
 
          (*This message models possible leaks of session keys.  The nonces
            identify the protocol run.  Quoting Server here ensures they are
            correct. *)
-    Oops "[| evso: yahalom;  
+    Oops "[| evso \\<in> yahalom;  
              Says Server A {|Nonce NB, 
                              Crypt (shrK A) {|Agent B, Key K, Nonce NA|},
-                             X|}  : set evso |]
-          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso : yahalom"
+                             X|}  \\<in> set evso |]
+          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \\<in> yahalom"
 
 end
--- a/src/HOL/Auth/Yahalom_Bad.ML	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Yahalom_Bad.ML	Tue Feb 27 16:13:23 2001 +0100
@@ -11,9 +11,9 @@
 *)
 
 (*A "possibility property": there are traces that reach the end*)
-Goal "A ~= Server \
-\     ==> EX X NB K. EX evs: yahalom.          \
-\            Says A B {|X, Crypt K (Nonce NB)|} : set evs";
+Goal "A \\<noteq> Server \
+\     ==> \\<exists>X NB K. \\<exists>evs \\<in> yahalom.          \
+\            Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (yahalom.Nil RS 
           yahalom.YM1 RS yahalom.Reception RS
@@ -22,21 +22,18 @@
 by possibility_tac;
 result();
 
-Goal "[| Gets B X : set evs; evs : yahalom |] ==> EX A. Says A B X : set evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> yahalom |] ==> \\<exists>A. Says A B X \\<in> set evs";
 by (etac rev_mp 1);
 by (etac yahalom.induct 1);
 by Auto_tac;
 qed "Gets_imp_Says";
 
 (*Must be proved separately for each protocol*)
-Goal "[| Gets B X : set evs; evs : yahalom |]  ==> X : knows Spy evs";
+Goal "[| Gets B X \\<in> set evs; evs \\<in> yahalom |]  ==> X \\<in> knows Spy evs";
 by (blast_tac (claset() addSDs [Gets_imp_Says, Says_imp_knows_Spy]) 1);
 qed"Gets_imp_knows_Spy";
 AddDs [Gets_imp_knows_Spy RS parts.Inj];
 
-fun g_not_bad_tac s = 
-  ftac Gets_imp_Says THEN' assume_tac THEN' not_bad_tac s;
-
 
 (**** Inductive proofs about yahalom ****)
 
@@ -44,22 +41,22 @@
 (** For reasoning about the encrypted portion of messages **)
 
 (*Lets us treat YM4 using a similar argument as for the Fake case.*)
-Goal "[| Gets A {|Crypt (shrK A) Y, X|} : set evs;  evs : yahalom |]  \
-\     ==> X : analz (knows Spy evs)";
+Goal "[| Gets A {|Crypt (shrK A) Y, X|} \\<in> set evs;  evs \\<in> yahalom |]  \
+\     ==> X \\<in> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [Gets_imp_knows_Spy RS analz.Inj]) 1);
 qed "YM4_analz_knows_Spy";
 
 bind_thm ("YM4_parts_knows_Spy",
           YM4_analz_knows_Spy RS (impOfSubs analz_subset_parts));
 
-(*For proving the easier theorems about X ~: parts (knows Spy evs).*)
+(*For proving the easier theorems about X \\<notin> parts (knows Spy evs).*)
 fun parts_knows_Spy_tac i = 
   EVERY
    [ftac YM4_parts_knows_Spy (i+6), assume_tac (i+6),
     prove_simple_subgoals_tac i];
 
 (*Induction for regularity theorems.  If induction formula has the form
-   X ~: analz (knows Spy evs) --> ... then it shortens the proof by discarding
+   X \\<notin> analz (knows Spy evs) --> ... then it shortens the proof by discarding
    needless information about analz (insert X (knows Spy evs))  *)
 fun parts_induct_tac i = 
     etac yahalom.induct i
@@ -68,18 +65,18 @@
     THEN  parts_knows_Spy_tac i;
 
 
-(** Theorems of the form X ~: parts (knows Spy evs) imply that NOBODY
+(** Theorems of the form X \\<notin> parts (knows Spy evs) imply that NOBODY
     sends messages containing X! **)
 
 (*Spy never sees another agent's shared key! (unless it's bad at start)*)
-Goal "evs : yahalom ==> (Key (shrK A) : parts (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> yahalom ==> (Key (shrK A) \\<in> parts (knows Spy evs)) = (A \\<in> bad)";
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 by (ALLGOALS Blast_tac);
 qed "Spy_see_shrK";
 Addsimps [Spy_see_shrK];
 
-Goal "evs : yahalom ==> (Key (shrK A) : analz (knows Spy evs)) = (A : bad)";
+Goal "evs \\<in> yahalom ==> (Key (shrK A) \\<in> analz (knows Spy evs)) = (A \\<in> bad)";
 by (auto_tac(claset() addDs [impOfSubs analz_subset_parts], simpset()));
 qed "Spy_analz_shrK";
 Addsimps [Spy_analz_shrK];
@@ -89,20 +86,15 @@
 
 
 (*Nobody can have used non-existent keys!  Needed to apply analz_insert_Key*)
-Goal "evs : yahalom ==>          \
-\      Key K ~: used evs --> K ~: keysFor (parts (knows Spy evs))";
+Goal "evs \\<in> yahalom ==>          \
+\      Key K \\<notin> used evs --> K \\<notin> keysFor (parts (knows Spy evs))";
 by (parts_induct_tac 1);
 (*Fake*)
 by (blast_tac (claset() addSDs [keysFor_parts_insert]) 1);
 (*YM2-4: Because Key K is not fresh, etc.*)
 by (REPEAT (blast_tac (claset() addSEs knows_Spy_partsEs) 1));
 qed_spec_mp "new_keys_not_used";
-
-bind_thm ("new_keys_not_analzd",
-          [analz_subset_parts RS keysFor_mono,
-           new_keys_not_used] MRS contra_subsetD);
-
-Addsimps [new_keys_not_used, new_keys_not_analzd];
+Addsimps [new_keys_not_used];
 
 
 (*For proofs involving analz.*)
@@ -112,18 +104,18 @@
 (****
  The following is to prove theorems of the form
 
-  Key K : analz (insert (Key KAB) (knows Spy evs)) ==>
-  Key K : analz (knows Spy evs)
+  Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) ==>
+  Key K \\<in> analz (knows Spy evs)
 
  A more general formula must be proved inductively.
 ****)
 
 (** Session keys are not used to encrypt other session keys **)
 
-Goal "evs : yahalom ==>                              \
-\  ALL K KK. KK <= - (range shrK) -->                \
-\         (Key K : analz (Key`KK Un (knows Spy evs))) = \
-\         (K : KK | Key K : analz (knows Spy evs))";
+Goal "evs \\<in> yahalom ==>                              \
+\  \\<forall>K KK. KK <= - (range shrK) -->                \
+\         (Key K \\<in> analz (Key`KK Un (knows Spy evs))) = \
+\         (K \\<in> KK | Key K \\<in> analz (knows Spy evs))";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (REPEAT_FIRST (resolve_tac [allI, impI]));
@@ -133,9 +125,9 @@
 by (spy_analz_tac 1);
 qed_spec_mp "analz_image_freshK";
 
-Goal "[| evs : yahalom;  KAB ~: range shrK |]                  \
-\      ==> Key K : analz (insert (Key KAB) (knows Spy evs)) =  \
-\          (K = KAB | Key K : analz (knows Spy evs))";
+Goal "[| evs \\<in> yahalom;  KAB \\<notin> range shrK |]                  \
+\      ==> Key K \\<in> analz (insert (Key KAB) (knows Spy evs)) =  \
+\          (K = KAB | Key K \\<in> analz (knows Spy evs))";
 by (asm_simp_tac (analz_image_freshK_ss addsimps [analz_image_freshK]) 1);
 qed "analz_insert_freshK";
 
@@ -143,10 +135,10 @@
 (*** The Key K uniquely identifies the Server's  message. **)
 
 Goal "[| Says Server A                                                 \
-\         {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} : set evs;  \
+\         {|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} \\<in> set evs;  \
 \       Says Server A'                                                 \
-\         {|Crypt (shrK A') {|Agent B', Key K, na', nb'|}, X'|} : set evs; \
-\       evs : yahalom |]                                    \
+\         {|Crypt (shrK A') {|Agent B', Key K, na', nb'|}, X'|} \\<in> set evs; \
+\       evs \\<in> yahalom |]                                    \
 \    ==> A=A' & B=B' & na=na' & nb=nb'";
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -161,12 +153,12 @@
 
 (** Crucial secrecy property: Spy does not see the keys sent in msg YM3 **)
 
-Goal "[| A ~: bad;  B ~: bad;  evs : yahalom |]                \
+Goal "[| A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
 \     ==> Says Server A                                        \
 \           {|Crypt (shrK A) {|Agent B, Key K, na, nb|},       \
 \             Crypt (shrK B) {|Agent A, Key K|}|}              \
-\          : set evs -->                                       \
-\         Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs -->                                       \
+\         Key K \\<notin> analz (knows Spy evs)";
 by (etac yahalom.induct 1);
 by analz_knows_Spy_tac;
 by (ALLGOALS
@@ -186,9 +178,9 @@
 Goal "[| Says Server A                                         \
 \           {|Crypt (shrK A) {|Agent B, Key K, na, nb|},       \
 \             Crypt (shrK B) {|Agent A, Key K|}|}              \
-\          : set evs;                                          \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                \
-\     ==> Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs;                                          \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSEs [lemma]) 1);
 qed "Spy_not_see_encrypted_key";
 
@@ -196,21 +188,21 @@
 (** Security Guarantee for A upon receiving YM3 **)
 
 (*If the encrypted message appears then it originated with the Server*)
-Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} : parts (knows Spy evs); \
-\        A ~: bad;  evs : yahalom |]                          \
+Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \\<in> parts (knows Spy evs); \
+\        A \\<notin> bad;  evs \\<in> yahalom |]                          \
 \      ==> Says Server A                                      \
 \           {|Crypt (shrK A) {|Agent B, Key K, na, nb|},      \
 \             Crypt (shrK B) {|Agent A, Key K|}|}             \
-\          : set evs";
+\          \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
 qed "A_trusts_YM3";
 
 (*The obvious combination of A_trusts_YM3 with Spy_not_see_encrypted_key*)
-Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} : parts (knows Spy evs); \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                \
-\     ==> Key K ~: analz (knows Spy evs)";
+Goal "[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \\<in> parts (knows Spy evs); \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [A_trusts_YM3, Spy_not_see_encrypted_key]) 1);
 qed "A_gets_good_key";
 
@@ -218,13 +210,13 @@
 
 (*B knows, by the first part of A's message, that the Server distributed 
   the key for A and B.  But this part says nothing about nonces.*)
-Goal "[| Crypt (shrK B) {|Agent A, Key K|} : parts (knows Spy evs);  \
-\        B ~: bad;  evs : yahalom |]                                 \
-\     ==> EX NA NB. Says Server A                                    \
+Goal "[| Crypt (shrK B) {|Agent A, Key K|} \\<in> parts (knows Spy evs);  \
+\        B \\<notin> bad;  evs \\<in> yahalom |]                                 \
+\     ==> \\<exists>NA NB. Says Server A                                    \
 \                     {|Crypt (shrK A) {|Agent B, Key K,             \
 \                                        Nonce NA, Nonce NB|},       \
 \                       Crypt (shrK B) {|Agent A, Key K|}|}          \
-\                    : set evs";
+\                    \\<in> set evs";
 by (etac rev_mp 1);
 by (parts_induct_tac 1);
 by (Fake_parts_insert_tac 1);
@@ -241,58 +233,53 @@
   the key quoting nonce NB.  This part says nothing about agent names. 
   Secrecy of K is assumed; the valid Yahalom proof uses (and later proves)
   the secrecy of NB.*)
-Goal "evs : yahalom                                          \
-\     ==> Key K ~: analz (knows Spy evs) -->                 \
-\         Crypt K (Nonce NB) : parts (knows Spy evs) -->     \
-\         (EX A B NA. Says Server A                          \
+Goal "evs \\<in> yahalom                                          \
+\     ==> Key K \\<notin> analz (knows Spy evs) -->                 \
+\         Crypt K (Nonce NB) \\<in> parts (knows Spy evs) -->     \
+\         (\\<exists>A B NA. Says Server A                          \
 \                     {|Crypt (shrK A) {|Agent B, Key K,     \
 \                               Nonce NA, Nonce NB|},        \
 \                       Crypt (shrK B) {|Agent A, Key K|}|}  \
-\                    : set evs)";
+\                    \\<in> set evs)";
 by (parts_induct_tac 1);
 by (ALLGOALS Clarify_tac);
 (*YM3 & Fake*)
 by (Blast_tac 2);
 by (Fake_parts_insert_tac 1);
 (*YM4*)
-(*A is uncompromised because NB is secure*)
-by (g_not_bad_tac "A" 1);
-(*A's certificate guarantees the existence of the Server message*)
-by (blast_tac (claset() addDs [Says_imp_knows_Spy RS parts.Inj RS parts.Fst RS
-			       A_trusts_YM3]) 1);
+(*A is uncompromised because NB is secure;
+  A's certificate guarantees the existence of the Server message*)
+by (blast_tac (claset() addSDs [Gets_imp_Says, Crypt_Spy_analz_bad]
+			addDs  [Says_imp_spies, analz.Inj, 
+			        parts.Inj RS parts.Fst RS A_trusts_YM3]) 1);
 bind_thm ("B_trusts_YM4_newK", result() RS mp RSN (2, rev_mp));
 
 
 (*B's session key guarantee from YM4.  The two certificates contribute to a
   single conclusion about the Server's message. *)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},                    \
-\                 Crypt K (Nonce NB)|} : set evs;                       \
+\                 Crypt K (Nonce NB)|} \\<in> set evs;                       \
 \        Says B Server                                                  \
 \          {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}  \
-\          : set evs;                                                   \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                         \
-\      ==> EX na nb. Says Server A                                      \
+\          \\<in> set evs;                                                   \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                         \
+\      ==> \\<exists>na nb. Says Server A                                      \
 \                  {|Crypt (shrK A) {|Agent B, Key K, na, nb|},         \
 \                    Crypt (shrK B) {|Agent A, Key K|}|}                \
-\            : set evs";
-by (etac (Gets_imp_knows_Spy RS parts.Inj RS MPair_parts) 1 THEN
-    assume_tac 1 THEN dtac B_trusts_YM4_shrK 1);
-by (dtac B_trusts_YM4_newK 3);
-by (REPEAT_FIRST (eresolve_tac [asm_rl, exE]));
-by (etac Spy_not_see_encrypted_key 1 THEN REPEAT (assume_tac 1));
-by (ftac unique_session_keys 1 THEN REPEAT (assume_tac 1));
-by (blast_tac (claset() addDs []) 1);
+\            \\<in> set evs";
+by (blast_tac (claset() addDs [B_trusts_YM4_newK, B_trusts_YM4_shrK, 
+                          Spy_not_see_encrypted_key, unique_session_keys]) 1);
 qed "B_trusts_YM4";
 
 
 (*The obvious combination of B_trusts_YM4 with Spy_not_see_encrypted_key*)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},                   \
-\                    Crypt K (Nonce NB)|} : set evs;                   \
+\                    Crypt K (Nonce NB)|} \\<in> set evs;                   \
 \        Says B Server                                                 \
 \          {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|} \
-\          : set evs;                                                  \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]                \
-\     ==> Key K ~: analz (knows Spy evs)";
+\          \\<in> set evs;                                                  \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]                \
+\     ==> Key K \\<notin> analz (knows Spy evs)";
 by (blast_tac (claset() addSDs [B_trusts_YM4, Spy_not_see_encrypted_key]) 1);
 qed "B_gets_good_key";
 
@@ -306,42 +293,37 @@
 (*Assuming the session key is secure, if both certificates are present then
   A has said NB.  We can't be sure about the rest of A's message, but only
   NB matters for freshness.*)  
-Goal "evs : yahalom                                              \
-\     ==> Key K ~: analz (knows Spy evs) -->                     \
-\         Crypt K (Nonce NB) : parts (knows Spy evs) -->         \
-\         Crypt (shrK B) {|Agent A, Key K|} : parts (knows Spy evs) --> \
-\         B ~: bad -->                                           \
-\         (EX X. Says A B {|X, Crypt K (Nonce NB)|} : set evs)";
+Goal "evs \\<in> yahalom                                              \
+\     ==> Key K \\<notin> analz (knows Spy evs) -->                     \
+\         Crypt K (Nonce NB) \\<in> parts (knows Spy evs) -->         \
+\         Crypt (shrK B) {|Agent A, Key K|} \\<in> parts (knows Spy evs) --> \
+\         B \\<notin> bad -->                                           \
+\         (\\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs)";
 by (parts_induct_tac 1);
 (*Fake*)
 by (Fake_parts_insert_tac 1);
 (*YM3: by new_keys_not_used we note that Crypt K (Nonce NB) could not exist*)
-by (fast_tac (claset() addSDs [Crypt_imp_keysFor] addss (simpset())) 1); 
+by (force_tac (claset() addSDs [Crypt_imp_keysFor], simpset()) 1); 
 (*YM4: was Crypt K (Nonce NB) the very last message?  If not, use ind. hyp.*)
 by (asm_simp_tac (simpset() addsimps [ex_disj_distrib]) 1);
 (*yes: apply unicity of session keys*)
-by (g_not_bad_tac "Aa" 1);
-by (blast_tac (claset() addSEs [MPair_parts]
-                        addSDs [A_trusts_YM3, B_trusts_YM4_shrK]
-		        addDs  [Says_imp_knows_Spy RS parts.Inj,
-				unique_session_keys]) 1);
+by (blast_tac (claset() addSDs [Gets_imp_Says, A_trusts_YM3, B_trusts_YM4_shrK,
+                                Crypt_Spy_analz_bad]
+		addDs  [Says_imp_knows_Spy RS parts.Inj, 
+                        Says_imp_spies RS analz.Inj, unique_session_keys]) 1);
 qed_spec_mp "A_Said_YM3_lemma";
 
 (*If B receives YM4 then A has used nonce NB (and therefore is alive).
   Moreover, A associates K with NB (thus is talking about the same run).
   Other premises guarantee secrecy of K.*)
 Goal "[| Gets B {|Crypt (shrK B) {|Agent A, Key K|},                   \
-\                 Crypt K (Nonce NB)|} : set evs;                      \
+\                 Crypt K (Nonce NB)|} \\<in> set evs;                      \
 \        Says B Server                                                 \
 \          {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|} \
-\          : set evs;                                                  \
-\        A ~: bad;  B ~: bad;  evs : yahalom |]       \
-\     ==> EX X. Says A B {|X, Crypt K (Nonce NB)|} : set evs";
-by (ftac B_trusts_YM4 1);
-by (REPEAT_FIRST assume_tac);
-by (etac (Gets_imp_knows_Spy RS parts.Inj RS MPair_parts) 1 THEN assume_tac 1);
-by (Clarify_tac 1);
-by (rtac A_Said_YM3_lemma 1);
-by (rtac Spy_not_see_encrypted_key 2);
-by (REPEAT_FIRST assume_tac);
+\          \\<in> set evs;                                                  \
+\        A \\<notin> bad;  B \\<notin> bad;  evs \\<in> yahalom |]       \
+\     ==> \\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \\<in> set evs";
+by (blast_tac (claset() addSIs [A_Said_YM3_lemma]
+                   addDs [Spy_not_see_encrypted_key, B_trusts_YM4,
+                          Gets_imp_Says, Says_imp_knows_Spy RS parts.Inj]) 1);
 qed_spec_mp "YM4_imp_A_Said_YM3";
--- a/src/HOL/Auth/Yahalom_Bad.thy	Tue Feb 27 12:28:42 2001 +0100
+++ b/src/HOL/Auth/Yahalom_Bad.thy	Tue Feb 27 16:13:23 2001 +0100
@@ -15,49 +15,49 @@
 inductive "yahalom"
   intrs 
          (*Initial trace is empty*)
-    Nil  "[]: yahalom"
+    Nil  "[] : yahalom"
 
          (*The spy MAY say anything he CAN say.  We do not expect him to
            invent new nonces here, but he can also use NS1.  Common to
            all similar protocols.*)
-    Fake "[| evs: yahalom;  X: synth (analz (knows Spy evs)) |]
-          ==> Says Spy B X  # evs : yahalom"
+    Fake "[| evsf \\<in> yahalom;  X \\<in> synth (analz (knows Spy evsf)) |]
+          ==> Says Spy B X  # evsf \\<in> yahalom"
 
          (*A message that has been sent can be received by the
            intended recipient.*)
-    Reception "[| evsr: yahalom;  Says A B X : set evsr |]
-               ==> Gets B X # evsr : yahalom"
+    Reception "[| evsr \\<in> yahalom;  Says A B X \\<in> set evsr |]
+               ==> Gets B X # evsr \\<in> yahalom"
 
          (*Alice initiates a protocol run*)
-    YM1  "[| evs1: yahalom;  Nonce NA ~: used evs1 |]
-          ==> Says A B {|Agent A, Nonce NA|} # evs1 : yahalom"
+    YM1  "[| evs1 \\<in> yahalom;  Nonce NA \\<notin> used evs1 |]
+          ==> Says A B {|Agent A, Nonce NA|} # evs1 \\<in> yahalom"
 
          (*Bob's response to Alice's message.*)
-    YM2  "[| evs2: yahalom;  Nonce NB ~: used evs2;
-             Gets B {|Agent A, Nonce NA|} : set evs2 |]
+    YM2  "[| evs2 \\<in> yahalom;  Nonce NB \\<notin> used evs2;
+             Gets B {|Agent A, Nonce NA|} \\<in> set evs2 |]
           ==> Says B Server 
                   {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
-                # evs2 : yahalom"
+                # evs2 \\<in> yahalom"
 
          (*The Server receives Bob's message.  He responds by sending a
             new session key to Alice, with a packet for forwarding to Bob.*)
-    YM3  "[| evs3: yahalom;  Key KAB ~: used evs3;
+    YM3  "[| evs3 \\<in> yahalom;  Key KAB \\<notin> used evs3;
              Gets Server 
                   {|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
-               : set evs3 |]
+               \\<in> set evs3 |]
           ==> Says Server A
                    {|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|},
                      Crypt (shrK B) {|Agent A, Key KAB|}|}
-                # evs3 : yahalom"
+                # evs3 \\<in> yahalom"
 
          (*Alice receives the Server's (?) message, checks her Nonce, and
            uses the new session key to send Bob his Nonce.  The premise
-           A ~= Server is needed to prove Says_Server_not_range.*)
-    YM4  "[| evs4: yahalom;  A ~= Server;
+           A \\<noteq> Server is needed to prove Says_Server_not_range.*)
+    YM4  "[| evs4 \\<in> yahalom;  A \\<noteq> Server;
              Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
-                : set evs4;
-             Says A B {|Agent A, Nonce NA|} : set evs4 |]
-          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 : yahalom"
+                \\<in> set evs4;
+             Says A B {|Agent A, Nonce NA|} \\<in> set evs4 |]
+          ==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 \\<in> yahalom"
 
          (*This message models possible leaks of session keys.  The Nonces
            identify the protocol run.  Quoting Server here ensures they are