Moving common declarations and proofs from theories "Shared"
authorpaulson
Fri, 11 Jul 1997 13:26:15 +0200
changeset 3512 9dcb4daa15e8
parent 3511 da4dd8b7ced4
child 3513 4d4f8c18255e
Moving common declarations and proofs from theories "Shared" and "Public" to "Event". NB the original "Event" theory was later renamed "Shared". Addition of the Notes constructor to datatype "event".
src/HOL/Auth/Event.ML
src/HOL/Auth/Event.thy
src/HOL/Auth/Public.ML
src/HOL/Auth/Public.thy
src/HOL/Auth/Shared.ML
src/HOL/Auth/Shared.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/Event.ML	Fri Jul 11 13:26:15 1997 +0200
@@ -0,0 +1,152 @@
+(*  Title:      HOL/Auth/Event
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1996  University of Cambridge
+
+Theory of events for security protocols
+
+Datatype of events; function "sees"; freshness
+*)
+
+open Event;
+
+(*** Function "sees" ***)
+
+(** Specialized rewrite rules for (sees lost A (Says...#evs)) **)
+
+goal thy "sees lost B (Says A B X # evs) = insert X (sees lost B evs)";
+by (Simp_tac 1);
+qed "sees_own";
+
+goal thy "sees lost B (Notes A X # evs) = \
+\         (if A=B then insert X (sees lost B evs) else sees lost B evs)";
+by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
+qed "sees_Notes";
+
+(** Three special-case rules for rewriting of sees lost A **)
+
+goal thy "!!A. Server ~= B ==> \
+\          sees lost Server (Says A B X # evs) = sees lost Server evs";
+by (Asm_simp_tac 1);
+qed "sees_Server";
+
+goal thy "!!A. Friend i ~= B ==> \
+\          sees lost (Friend i) (Says A B X # evs) = sees lost (Friend i) evs";
+by (Asm_simp_tac 1);
+qed "sees_Friend";
+
+goal thy "sees lost Spy (Says A B X # evs) = insert X (sees lost Spy evs)";
+by (Simp_tac 1);
+qed "sees_Spy";
+
+goal thy "sees lost A (Says A' B X # evs) <= insert X (sees lost A evs)";
+by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
+by (Fast_tac 1);
+qed "sees_Says_subset_insert";
+
+goal thy "sees lost A evs <= sees lost A (Says A' B X # evs)";
+by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
+by (Fast_tac 1);
+qed "sees_subset_sees_Says";
+
+goal thy "sees lost A evs <= sees lost A (Notes A' X # evs)";
+by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
+by (Fast_tac 1);
+qed "sees_subset_sees_Notes";
+
+(*Pushing Unions into parts.  One of the agents A is B, and thus sees Y.*)
+goal thy "(UN A. parts (sees lost A (Says B C Y # evs))) = \
+\         parts {Y} Un (UN A. parts (sees lost A evs))";
+by (Step_tac 1);
+by (etac rev_mp 1);     (*split_tac does not work on assumptions*)
+by (ALLGOALS
+    (fast_tac (!claset addss (!simpset addsimps [parts_Un] 
+				       setloop split_tac [expand_if]))));
+qed "UN_parts_sees_Says";
+
+goal thy "(UN A. parts (sees lost A (Notes B Y # evs))) = \
+\         parts {Y} Un (UN A. parts (sees lost A evs))";
+by (Step_tac 1);
+by (etac rev_mp 1);     (*split_tac does not work on assumptions*)
+by (ALLGOALS
+    (fast_tac (!claset addss (!simpset addsimps [parts_Un] 
+				       setloop split_tac [expand_if]))));
+qed "UN_parts_sees_Notes";
+
+goal thy "Says A B X : set evs --> X : sees lost Spy evs";
+by (list.induct_tac "evs" 1);
+by (Auto_tac ());
+qed_spec_mp "Says_imp_sees_Spy";
+
+(*Use with addSEs to derive contradictions from old Says events containing
+  items known to be fresh*)
+val sees_Spy_partsEs = make_elim (Says_imp_sees_Spy RS parts.Inj):: partsEs;
+
+Addsimps [sees_own, sees_Notes, sees_Server, sees_Friend, sees_Spy];
+
+(**** NOTE REMOVAL--laws above are cleaner--def of sees1 is messy ****)
+Delsimps [sees_Cons];   
+
+
+(*** Fresh nonces ***)
+
+goalw thy [used_def] "!!X. X: parts (sees lost B evs) ==> X: used evs";
+by (etac (impOfSubs parts_mono) 1);
+by (Fast_tac 1);
+qed "usedI";
+AddIs [usedI];
+
+goal thy "used (Says A B X # evs) = parts{X} Un used evs";
+by (simp_tac (!simpset addsimps [used_def, UN_parts_sees_Says]) 1);
+qed "used_Says";
+Addsimps [used_Says];
+
+goal thy "used (Notes A X # evs) = parts{X} Un used evs";
+by (simp_tac (!simpset delsimps [sees_Notes]
+		       addsimps [used_def, UN_parts_sees_Notes]) 1);
+qed "used_Notes";
+Addsimps [used_Notes];
+
+(*These two facts about "used" are unused.*)
+goal thy "used [] <= used l";
+by (list.induct_tac "l" 1);
+by (event.induct_tac "a" 2);
+by (ALLGOALS Asm_simp_tac);
+by (ALLGOALS Blast_tac);
+qed "used_nil_subset";
+
+goal thy "used l <= used (l@l')";
+by (list.induct_tac "l" 1);
+by (simp_tac (!simpset addsimps [used_nil_subset]) 1);
+by (event.induct_tac "a" 1);
+by (ALLGOALS Asm_simp_tac);
+by (ALLGOALS Blast_tac);
+qed "used_subset_append";
+
+
+(** Simplifying   parts (insert X (sees lost A evs))
+      = parts {X} Un parts (sees lost A evs) -- since general case loops*)
+
+val parts_insert_sees = 
+    parts_insert |> read_instantiate_sg (sign_of thy)
+                                        [("H", "sees lost A evs")]
+                 |> standard;
+
+
+
+(*For proving theorems of the form X ~: analz (sees Spy evs) --> P
+  New events added by induction to "evs" are discarded.  Provided 
+  this information isn't needed, the proof will be much shorter, since
+  it will omit complicated reasoning about analz.*)
+val analz_mono_contra_tac = 
+  let val impI' = read_instantiate_sg (sign_of thy)
+                [("P", "?Y ~: analz (sees lost ?A ?evs)")] impI;
+  in
+    rtac impI THEN' 
+    REPEAT1 o 
+      (dresolve_tac 
+       [sees_subset_sees_Says  RS analz_mono RS contra_subsetD,
+	sees_subset_sees_Notes RS analz_mono RS contra_subsetD])
+    THEN'
+    mp_tac
+  end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Auth/Event.thy	Fri Jul 11 13:26:15 1997 +0200
@@ -0,0 +1,42 @@
+(*  Title:      HOL/Auth/Event
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1996  University of Cambridge
+
+Theory of events for security protocols
+
+Datatype of events; function "sees"; freshness
+*)
+
+Event = Message + List + 
+
+consts  (*Initial states of agents -- parameter of the construction*)
+  initState :: [agent set, agent] => msg set
+
+datatype  (*Messages--could add another constructor for agent knowledge*)
+  event = Says  agent agent msg
+        | Notes agent       msg
+
+consts  
+  sees1 :: [agent, event] => msg set
+
+primrec sees1 event
+           (*Spy reads all traffic whether addressed to him or not*)
+  sees1_Says  "sees1 A (Says A' B X)  = (if A:{B,Spy} then {X} else {})"
+  sees1_Notes "sees1 A (Notes A' X)   = (if A = A'    then {X} else {})"
+
+consts  
+  sees :: [agent set, agent, event list] => msg set
+
+primrec sees list
+  sees_Nil  "sees lost A []       = initState lost A"
+  sees_Cons "sees lost A (ev#evs) = sees1 A ev Un sees lost A evs"
+
+
+constdefs
+  (*Set of items that might be visible to somebody: complement of the set
+        of fresh items*)
+  used :: event list => msg set
+    "used evs == parts (UN lost B. sees lost B evs)"
+
+end
--- a/src/HOL/Auth/Public.ML	Wed Jul 09 17:00:34 1997 +0200
+++ b/src/HOL/Auth/Public.ML	Fri Jul 11 13:26:15 1997 +0200
@@ -11,41 +11,6 @@
 
 open Public;
 
-(*Holds because Friend is injective: thus cannot prove for all f*)
-goal thy "(Friend x : Friend``A) = (x:A)";
-by (Auto_tac());
-qed "Friend_image_eq";
-Addsimps [Friend_image_eq];
-
-Addsimps [Un_insert_left, Un_insert_right];
-
-(*By default only o_apply is built-in.  But in the presence of eta-expansion
-  this means that some terms displayed as (f o g) will be rewritten, and others
-  will not!*)
-Addsimps [o_def];
-
-goalw thy [keysFor_def] "keysFor (parts (initState lost C)) = {}";
-by (agent.induct_tac "C" 1);
-by (auto_tac (!claset addIs [range_eqI], !simpset));
-qed "keysFor_parts_initState";
-Addsimps [keysFor_parts_initState];
-
-goalw thy [keysFor_def] "keysFor (Key``E) = {}";
-by (Auto_tac ());
-qed "keysFor_image_Key";
-Addsimps [keysFor_image_Key];
-
-
-(*** Function "sees" ***)
-
-goal thy
-    "!!evs. lost' <= lost ==> sees lost' A evs <= sees lost A evs";
-by (list.induct_tac "evs" 1);
-by (agent.induct_tac "A" 1);
-by (event.induct_tac "a" 2);
-by (Auto_tac ());
-qed "sees_mono";
-
 (*** Basic properties of pubK & priK ***)
 
 AddIffs [inj_pubK RS inj_eq];
@@ -69,81 +34,44 @@
 
 AddIffs [not_isSymKey_pubK, not_isSymKey_priK];
 
+(** Rewrites should not refer to  initState(Friend i) 
+    -- not in normal form! **)
+
+goalw thy [keysFor_def] "keysFor (parts (initState lost C)) = {}";
+by (agent.induct_tac "C" 1);
+by (auto_tac (!claset addIs [range_eqI], !simpset));
+qed "keysFor_parts_initState";
+Addsimps [keysFor_parts_initState];
+
+
+(*** Function "sees" ***)
 
 (*Agents see their own private keys!*)
 goal thy "A ~= Spy --> Key (priK A) : sees lost A evs";
 by (list.induct_tac "evs" 1);
 by (agent.induct_tac "A" 1);
-by (Auto_tac ());
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
 qed_spec_mp "sees_own_priK";
 
 (*All public keys are visible to all*)
 goal thy "Key (pubK A) : sees lost B evs";
 by (list.induct_tac "evs" 1);
 by (agent.induct_tac "B" 1);
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
 by (Auto_tac ());
 qed_spec_mp "sees_pubK";
 
 (*Spy sees private keys of lost agents!*)
 goal thy "!!A. A: lost ==> Key (priK A) : sees lost Spy evs";
 by (list.induct_tac "evs" 1);
-by (Auto_tac());
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
+by (Blast_tac 1);
 qed "Spy_sees_lost";
 
 AddIffs [sees_pubK, sees_pubK RS analz.Inj];
 AddSIs  [sees_own_priK, Spy_sees_lost];
 
 
-(** Specialized rewrite rules for (sees lost A (Says...#evs)) **)
-
-goal thy "sees lost B (Says A B X # evs) = insert X (sees lost B evs)";
-by (Simp_tac 1);
-qed "sees_own";
-
-goal thy "!!A. Server ~= B ==> \
-\          sees lost Server (Says A B X # evs) = sees lost Server evs";
-by (Asm_simp_tac 1);
-qed "sees_Server";
-
-goal thy "!!A. Friend i ~= B ==> \
-\          sees lost (Friend i) (Says A B X # evs) = sees lost (Friend i) evs";
-by (Asm_simp_tac 1);
-qed "sees_Friend";
-
-goal thy "sees lost Spy (Says A B X # evs) = insert X (sees lost Spy evs)";
-by (Simp_tac 1);
-qed "sees_Spy";
-
-goal thy "sees lost A (Says A' B X # evs) <= insert X (sees lost A evs)";
-by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
-by (Fast_tac 1);
-qed "sees_Says_subset_insert";
-
-goal thy "sees lost A evs <= sees lost A (Says A' B X # evs)";
-by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
-by (Fast_tac 1);
-qed "sees_subset_sees_Says";
-
-(*Pushing Unions into parts.  One of the agents A is B, and thus sees Y.
-  Once used to prove new_keys_not_seen; now obsolete.*)
-goal thy "(UN A. parts (sees lost A (Says B C Y # evs))) = \
-\         parts {Y} Un (UN A. parts (sees lost A evs))";
-by (Step_tac 1);
-by (etac rev_mp 1);     (*split_tac does not work on assumptions*)
-by (ALLGOALS
-    (fast_tac (!claset addss (!simpset addsimps [parts_Un, sees_Cons] 
-				            setloop split_tac [expand_if]))));
-qed "UN_parts_sees_Says";
-
-goal thy "Says A B X : set evs --> X : sees lost Spy evs";
-by (list.induct_tac "evs" 1);
-by (Auto_tac ());
-qed_spec_mp "Says_imp_sees_Spy";
-
-(*Use with addSEs to derive contradictions from old Says events containing
-  items known to be fresh*)
-val sees_Spy_partsEs = make_elim (Says_imp_sees_Spy RS parts.Inj):: partsEs;
-
 (*For not_lost_tac*)
 goal thy "!!A. [| Crypt (pubK A) X : analz (sees lost Spy evs);  A: lost |] \
 \              ==> X : analz (sees lost Spy evs)";
@@ -162,9 +90,6 @@
          (has_fewer_prems 1, size_of_thm)
          (Step_tac 1));
 
-Addsimps [sees_own, sees_Server, sees_Friend, sees_Spy];
-Delsimps [sees_Cons];   (**** NOTE REMOVAL -- laws above are cleaner ****)
-
 
 (*** Fresh nonces ***)
 
@@ -172,17 +97,15 @@
 by (agent.induct_tac "B" 1);
 by (Auto_tac ());
 qed "Nonce_notin_initState";
-
 AddIffs [Nonce_notin_initState];
 
-goalw thy [used_def] "!!X. X: parts (sees lost B evs) ==> X: used evs";
-by (etac (impOfSubs parts_mono) 1);
-by (Fast_tac 1);
-qed "usedI";
+goal thy "Nonce N ~: used []";
+by (simp_tac (!simpset addsimps [used_def]) 1);
+qed "Nonce_notin_used_empty";
+Addsimps [Nonce_notin_used_empty];
 
-AddIs [usedI];
 
-(** A supply of fresh nonces for possibility theorems. **)
+(*** Supply fresh nonces for possibility theorems. ***)
 
 goalw thy [used_def] "EX N. ALL n. N<=n --> Nonce n ~: used evs";
 by (list.induct_tac "evs" 1);
@@ -191,16 +114,12 @@
 by (Full_simp_tac 1);
 (*Inductive step*)
 by (event.induct_tac "a" 1);
-by (full_simp_tac (!simpset addsimps [UN_parts_sees_Says]) 1);
-by (msg.induct_tac "msg" 1);
-by (ALLGOALS (asm_simp_tac (!simpset addsimps [exI, parts_insert2])));
-by (Step_tac 1);
-(*MPair case*)
-by (res_inst_tac [("x","Na+Nb")] exI 2);
-by (fast_tac (!claset addSEs [add_leE]) 2);
-(*Nonce case*)
-by (res_inst_tac [("x","N + Suc nat")] exI 1);
-by (fast_tac (!claset addSEs [add_leE] addaltern trans_tac) 1);
+by (ALLGOALS (full_simp_tac 
+	      (!simpset delsimps [sees_Notes]
+		        addsimps [UN_parts_sees_Says,
+				  UN_parts_sees_Notes])));
+by (ALLGOALS (rtac (msg_Nonce_supply RS exE)));
+by (ALLGOALS (blast_tac (!claset addSEs [add_leE])));
 val lemma = result();
 
 goal thy "Nonce (@ N. Nonce N ~: used evs) ~: used evs";
@@ -211,35 +130,30 @@
 
 (*Tactic for possibility theorems*)
 val possibility_tac =
-    REPEAT 
-    (ALLGOALS (simp_tac (!simpset setSolver safe_solver))
+    REPEAT (*omit used_Says so that Nonces, Keys start from different traces!*)
+    (ALLGOALS (simp_tac (!simpset delsimps [used_Says] setSolver safe_solver))
      THEN
      REPEAT_FIRST (eq_assume_tac ORELSE' 
                    resolve_tac [refl, conjI, Nonce_supply]));
 
-(** Power of the Spy **)
+
+(*** Specialized rewriting for the analz_image_... theorems ***)
 
-(*The Spy can see more than anybody else, except for their initial state*)
-goal thy "sees lost A evs <= initState lost A Un sees lost Spy evs";
-by (list.induct_tac "evs" 1);
-by (event.induct_tac "a" 2);
-by (ALLGOALS (fast_tac (!claset addDs [sees_Says_subset_insert RS subsetD] 
-                                addss (!simpset))));
-qed "sees_agent_subset_sees_Spy";
+goal thy "insert (Key K) H = Key `` {K} Un H";
+by (Blast_tac 1);
+qed "insert_Key_singleton";
+
+goal thy "insert (Key K) (Key``KK Un C) = Key `` (insert K KK) Un C";
+by (Blast_tac 1);
+qed "insert_Key_image";
 
-(*The Spy can see more than anybody else who's lost their key!*)
-goal thy "A: lost --> A ~= Server --> sees lost A evs <= sees lost Spy evs";
-by (list.induct_tac "evs" 1);
-by (event.induct_tac "a" 2);
-by (agent.induct_tac "A" 1);
-by (auto_tac (!claset addDs [sees_Says_subset_insert RS subsetD], (!simpset)));
-qed_spec_mp "sees_lost_agent_subset_sees_Spy";
+(*Reverse the normal simplification of "image" to build up (not break down)
+  the set of keys.  Based on analz_image_freshK_ss, but simpler.*)
+val analz_image_keys_ss = 
+     !simpset delsimps [image_insert, image_Un]
+              addsimps [image_insert RS sym, image_Un RS sym,
+			rangeI, 
+			insert_Key_singleton, 
+			insert_Key_image, Un_assoc RS sym]
+              setloop split_tac [expand_if];
 
-
-(** Simplifying   parts (insert X (sees lost A evs))
-      = parts {X} Un parts (sees lost A evs) -- since general case loops*)
-
-val parts_insert_sees = 
-    parts_insert |> read_instantiate_sg (sign_of thy)
-                                        [("H", "sees lost A evs")]
-                 |> standard;
--- a/src/HOL/Auth/Public.thy	Wed Jul 09 17:00:34 1997 +0200
+++ b/src/HOL/Auth/Public.thy	Fri Jul 11 13:26:15 1997 +0200
@@ -3,12 +3,12 @@
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
 
-Theory of Public Keys (common to all symmetric-key protocols)
+Theory of Public Keys (common to all public-key protocols)
 
-Server keys; initial states of agents; new nonces and keys; function "sees" 
+Private and public keys; initial states of agents
 *)
 
-Public = Message + List + 
+Public = Event + 
 
 consts
   pubK    :: agent => key
@@ -19,9 +19,6 @@
 translations  (*BEWARE! expressions like  (inj priK)  will NOT work!*)
   "priK x"  == "invKey(pubK x)"
 
-consts  (*Initial states of agents -- parameter of the construction*)
-  initState :: [agent set, agent] => msg set
-
 primrec initState agent
         (*Agents know their private key and all public keys*)
   initState_Server  "initState lost Server     =    
@@ -32,31 +29,6 @@
  		         (Key``invKey``pubK``lost) Un (Key `` range pubK)"
 
 
-datatype  (*Messages, and components of agent stores*)
-  event = Says agent agent msg
-
-consts  
-  sees1 :: [agent, event] => msg set
-
-primrec sees1 event
-           (*Spy reads all traffic whether addressed to him or not*)
-  sees1_Says  "sees1 A (Says A' B X)  = (if A:{B,Spy} then {X} else {})"
-
-consts  
-  sees :: [agent set, agent, event list] => msg set
-
-primrec sees list
-  sees_Nil  "sees lost A []       = initState lost A"
-  sees_Cons "sees lost A (ev#evs) = sees1 A ev Un sees lost A evs"
-
-
-constdefs
-  (*Set of items that might be visible to somebody: complement of the set
-        of fresh items*)
-  used :: event list => msg set
-    "used evs == parts (UN lost B. sees lost B evs)"
-
-
 rules
   (*Public keys are disjoint, and never clash with private keys*)
   inj_pubK        "inj pubK"
--- a/src/HOL/Auth/Shared.ML	Wed Jul 09 17:00:34 1997 +0200
+++ b/src/HOL/Auth/Shared.ML	Fri Jul 11 13:26:15 1997 +0200
@@ -5,25 +5,12 @@
 
 Theory of Shared Keys (common to all symmetric-key protocols)
 
-Server keys; initial states of agents; freshness; function "sees" 
+Shared, long-term keys; initial states of agents
 *)
 
 
 open Shared;
 
-(*Holds because Friend is injective: thus cannot prove for all f*)
-goal thy "(Friend x : Friend``A) = (x:A)";
-by (Auto_tac());
-qed "Friend_image_eq";
-Addsimps [Friend_image_eq];
-
-Addsimps [Un_insert_left, Un_insert_right];
-
-(*By default only o_apply is built-in.  But in the presence of eta-expansion
-  this means that some terms displayed as (f o g) will be rewritten, and others
-  will not!*)
-Addsimps [o_def];
-
 (*** Basic properties of shrK ***)
 
 (*Injectiveness: Agents' long-term keys are distinct.*)
@@ -41,87 +28,26 @@
 qed "keysFor_parts_initState";
 Addsimps [keysFor_parts_initState];
 
-goalw thy [keysFor_def] "keysFor (Key``E) = {}";
-by (Auto_tac ());
-qed "keysFor_image_Key";
-Addsimps [keysFor_image_Key];
-
 
 (*** Function "sees" ***)
 
-goal thy
-    "!!evs. lost' <= lost ==> sees lost' A evs <= sees lost A evs";
-by (list.induct_tac "evs" 1);
-by (agent.induct_tac "A" 1);
-by (event.induct_tac "a" 2);
-by (Auto_tac ());
-qed "sees_mono";
-
 (*Agents see their own shared keys!*)
 goal thy "A ~= Spy --> Key (shrK A) : sees lost A evs";
 by (list.induct_tac "evs" 1);
 by (agent.induct_tac "A" 1);
-by (Auto_tac ());
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
+by (Blast_tac 1);
 qed_spec_mp "sees_own_shrK";
 
 (*Spy sees shared keys of lost agents!*)
 goal thy "!!A. A: lost ==> Key (shrK A) : sees lost Spy evs";
 by (list.induct_tac "evs" 1);
-by (Auto_tac());
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
+by (Blast_tac 1);
 qed "Spy_sees_lost";
 
 AddSIs [sees_own_shrK, Spy_sees_lost];
 
-(** Specialized rewrite rules for (sees lost A (Says...#evs)) **)
-
-goal thy "sees lost B (Says A B X # evs) = insert X (sees lost B evs)";
-by (Simp_tac 1);
-qed "sees_own";
-
-goal thy "!!A. Server ~= B ==> \
-\          sees lost Server (Says A B X # evs) = sees lost Server evs";
-by (Asm_simp_tac 1);
-qed "sees_Server";
-
-goal thy "!!A. Friend i ~= B ==> \
-\          sees lost (Friend i) (Says A B X # evs) = sees lost (Friend i) evs";
-by (Asm_simp_tac 1);
-qed "sees_Friend";
-
-goal thy "sees lost Spy (Says A B X # evs) = insert X (sees lost Spy evs)";
-by (Simp_tac 1);
-qed "sees_Spy";
-
-goal thy "sees lost A (Says A' B X # evs) <= insert X (sees lost A evs)";
-by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
-by (Blast_tac 1);
-qed "sees_Says_subset_insert";
-
-goal thy "sees lost A evs <= sees lost A (Says A' B X # evs)";
-by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
-by (Blast_tac 1);
-qed "sees_subset_sees_Says";
-
-(*Pushing Unions into parts.  One of the agents A is B, and thus sees Y.
-  Once used to prove new_keys_not_seen; now obsolete.*)
-goal thy "(UN A. parts (sees lost A (Says B C Y # evs))) = \
-\         parts {Y} Un (UN A. parts (sees lost A evs))";
-by (Step_tac 1);
-by (etac rev_mp 1);     (*split_tac does not work on assumptions*)
-by (ALLGOALS
-    (fast_tac (!claset addss (!simpset addsimps [parts_Un, sees_Cons] 
-				            setloop split_tac [expand_if]))));
-qed "UN_parts_sees_Says";
-
-goal thy "Says A B X : set evs --> X : sees lost Spy evs";
-by (list.induct_tac "evs" 1);
-by (Auto_tac ());
-qed_spec_mp "Says_imp_sees_Spy";
-
-(*Use with addSEs to derive contradictions from old Says events containing
-  items known to be fresh*)
-val sees_Spy_partsEs = make_elim (Says_imp_sees_Spy RS parts.Inj):: partsEs;
-
 (*For not_lost_tac*)
 goal thy "!!A. [| Crypt (shrK A) X : analz (sees lost Spy evs);  A: lost |] \
 \              ==> X : analz (sees lost Spy evs)";
@@ -140,25 +66,6 @@
          (has_fewer_prems 1, size_of_thm)
          (Step_tac 1));
 
-Addsimps [sees_own, sees_Server, sees_Friend, sees_Spy];
-Delsimps [sees_Cons];   (**** NOTE REMOVAL -- laws above are cleaner ****)
-
-
-(*** Fresh nonces ***)
-
-goal thy "Nonce N ~: parts (initState lost B)";
-by (agent.induct_tac "B" 1);
-by (Auto_tac ());
-qed "Nonce_notin_initState";
-
-AddIffs [Nonce_notin_initState];
-
-goalw thy [used_def] "!!X. X: parts (sees lost B evs) ==> X: used evs";
-by (etac (impOfSubs parts_mono) 1);
-by (Blast_tac 1);
-qed "usedI";
-
-AddIs [usedI];
 
 (** Fresh keys never clash with long-term shared keys **)
 
@@ -181,29 +88,23 @@
 Addsimps [Key_not_used, shrK_neq, shrK_neq RS not_sym];
 
 
-goal thy "used (Says A B X # evs) = parts{X} Un used evs";
-by (simp_tac (!simpset addsimps [used_def, UN_parts_sees_Says]) 1);
-qed "used_Says";
-Addsimps [used_Says];
+(*** Fresh nonces ***)
 
-goal thy "used [] <= used l";
-by (list.induct_tac "l" 1);
-by (event.induct_tac "a" 2);
-by (ALLGOALS Asm_simp_tac);
-by (Blast_tac 1);
-qed "used_nil_subset";
+goal thy "Nonce N ~: parts (initState lost B)";
+by (agent.induct_tac "B" 1);
+by (Auto_tac ());
+qed "Nonce_notin_initState";
+AddIffs [Nonce_notin_initState];
 
-goal thy "used l <= used (l@l')";
-by (list.induct_tac "l" 1);
-by (simp_tac (!simpset addsimps [used_nil_subset]) 1);
-by (event.induct_tac "a" 1);
-by (Asm_simp_tac 1);
-by (Blast_tac 1);
-qed "used_subset_append";
+goal thy "Nonce N ~: used []";
+by (simp_tac (!simpset addsimps [used_def]) 1);
+qed "Nonce_notin_used_empty";
+Addsimps [Nonce_notin_used_empty];
 
 
 (*** Supply fresh nonces for possibility theorems. ***)
 
+(*In any trace, there is an upper bound N on the greatest nonce in use.*)
 goalw thy [used_def] "EX N. ALL n. N<=n --> Nonce n ~: used evs";
 by (list.induct_tac "evs" 1);
 by (res_inst_tac [("x","0")] exI 1);
@@ -211,16 +112,12 @@
 by (Full_simp_tac 1);
 (*Inductive step*)
 by (event.induct_tac "a" 1);
-by (full_simp_tac (!simpset addsimps [UN_parts_sees_Says]) 1);
-by (msg.induct_tac "msg" 1);
-by (ALLGOALS (asm_simp_tac (!simpset addsimps [exI, parts_insert2])));
-by (Step_tac 1);
-(*MPair case*)
-by (res_inst_tac [("x","Na+Nb")] exI 2);
-by (blast_tac (!claset addSEs [add_leE]) 2);
-(*Nonce case*)
-by (res_inst_tac [("x","N + Suc nat")] exI 1);
-by (fast_tac (!claset addSEs [add_leE] addaltern trans_tac) 1);
+by (ALLGOALS (full_simp_tac 
+	      (!simpset delsimps [sees_Notes]
+		        addsimps [UN_parts_sees_Says,
+				  UN_parts_sees_Notes])));
+by (ALLGOALS (rtac (msg_Nonce_supply RS exE)));
+by (ALLGOALS (blast_tac (!claset addSEs [add_leE])));
 val lemma = result();
 
 goal thy "EX N. Nonce N ~: used evs";
@@ -302,8 +199,7 @@
 
 val possibility_tac =
     REPEAT (*omit used_Says so that Nonces, Keys start from different traces!*)
-    (ALLGOALS (simp_tac 
-               (!simpset delsimps [used_Says] setSolver safe_solver))
+    (ALLGOALS (simp_tac (!simpset delsimps [used_Says] setSolver safe_solver))
      THEN
      REPEAT_FIRST (eq_assume_tac ORELSE' 
                    resolve_tac [refl, conjI, Nonce_supply, Key_supply]));
@@ -317,35 +213,6 @@
      REPEAT_FIRST (resolve_tac [refl, conjI]));
 
 
-(** Power of the Spy **)
-
-(*The Spy can see more than anybody else, except for their initial state*)
-goal thy "sees lost A evs <= initState lost A Un sees lost Spy evs";
-by (list.induct_tac "evs" 1);
-by (event.induct_tac "a" 2);
-by (ALLGOALS Asm_simp_tac);
-by (ALLGOALS (blast_tac (!claset addDs [sees_Says_subset_insert RS subsetD])));
-qed "sees_agent_subset_sees_Spy";
-
-(*The Spy can see more than anybody else who's lost their key!*)
-goal thy "A: lost --> A ~= Server --> sees lost A evs <= sees lost Spy evs";
-by (list.induct_tac "evs" 1);
-by (event.induct_tac "a" 2);
-by (agent.induct_tac "A" 1);
-by (ALLGOALS Asm_simp_tac);
-by (ALLGOALS (blast_tac (!claset addDs [sees_Says_subset_insert RS subsetD])));
-qed_spec_mp "sees_lost_agent_subset_sees_Spy";
-
-
-(** Simplifying   parts (insert X (sees lost A evs))
-      = parts {X} Un parts (sees lost A evs) -- since general case loops*)
-
-val parts_insert_sees = 
-    parts_insert |> read_instantiate_sg (sign_of thy)
-                                        [("H", "sees lost A evs")]
-                 |> standard;
-
-
 (*** Specialized rewriting for analz_insert_freshK ***)
 
 goal thy "!!A. A <= Compl (range shrK) ==> shrK x ~: A";
--- a/src/HOL/Auth/Shared.thy	Wed Jul 09 17:00:34 1997 +0200
+++ b/src/HOL/Auth/Shared.thy	Fri Jul 11 13:26:15 1997 +0200
@@ -5,10 +5,10 @@
 
 Theory of Shared Keys (common to all symmetric-key protocols)
 
-Server keys; initial states of agents; new nonces and keys; function "sees" 
+Shared, long-term keys; initial states of agents
 *)
 
-Shared = Message + List + Finite +
+Shared = Event + Finite +
 
 consts
   shrK    :: agent => key  (*symmetric keys*)
@@ -17,9 +17,6 @@
   isSym_keys "isSymKey K"	(*All keys are symmetric*)
   inj_shrK   "inj shrK"		(*No two agents have the same long-term key*)
 
-consts  (*Initial states of agents -- parameter of the construction*)
-  initState :: [agent set, agent] => msg set
-
 primrec initState agent
         (*Server knows all long-term keys; other agents know only their own*)
   initState_Server  "initState lost Server     = Key `` range shrK"
@@ -27,31 +24,6 @@
   initState_Spy     "initState lost Spy        = Key``shrK``lost"
 
 
-datatype  (*Messages, and components of agent stores*)
-  event = Says agent agent msg
-
-consts  
-  sees1 :: [agent, event] => msg set
-
-primrec sees1 event
-           (*Spy reads all traffic whether addressed to him or not*)
-  sees1_Says  "sees1 A (Says A' B X)  = (if A:{B,Spy} then {X} else {})"
-
-consts  
-  sees :: [agent set, agent, event list] => msg set
-
-primrec sees list
-  sees_Nil  "sees lost A []       = initState lost A"
-  sees_Cons "sees lost A (ev#evs) = sees1 A ev Un sees lost A evs"
-
-
-constdefs
-  (*Set of items that might be visible to somebody: complement of the set
-        of fresh items*)
-  used :: event list => msg set
-    "used evs == parts (UN lost B. sees lost B evs)"
-
-
 rules
   (*Unlike the corresponding property of nonces, this cannot be proved.
     We have infinitely many agents and there is nothing to stop their