removed the infernal States, eqStates, compatible, etc.
authorpaulson
Mon, 01 Mar 1999 18:38:43 +0100
changeset 6295 351b3c2b0d83
parent 6294 bc084e1b4d8d
child 6296 9da8f9262c4c
removed the infernal States, eqStates, compatible, etc.
src/HOL/UNITY/Client.ML
src/HOL/UNITY/Client.thy
src/HOL/UNITY/Common.ML
src/HOL/UNITY/Comp.ML
src/HOL/UNITY/Comp.thy
src/HOL/UNITY/Constrains.ML
src/HOL/UNITY/Handshake.ML
src/HOL/UNITY/Handshake.thy
src/HOL/UNITY/Lift.thy
src/HOL/UNITY/Mutex.thy
src/HOL/UNITY/NSP_Bad.thy
src/HOL/UNITY/PPROD.ML
src/HOL/UNITY/PPROD.thy
src/HOL/UNITY/ROOT.ML
src/HOL/UNITY/Reach.thy
src/HOL/UNITY/SubstAx.ML
src/HOL/UNITY/Traces.ML
src/HOL/UNITY/Traces.thy
src/HOL/UNITY/UNITY.ML
src/HOL/UNITY/Union.ML
src/HOL/UNITY/Union.thy
src/HOL/UNITY/WFair.ML
--- a/src/HOL/UNITY/Client.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Client.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -90,8 +90,7 @@
 
 (*** Towards proving the liveness property ***)
 
-Goal "States Cli_prg = States G \
-\     ==> Cli_prg Join G   \
+Goal "Cli_prg Join G   \
 \       : transient {s. size (giv s) = Suc k &  \
 \                       size (rel s) = k & ask s ! k <= giv s ! k}";
 by (res_inst_tac [("act", "rel_act")] transient_mem 1);
@@ -131,13 +130,11 @@
 by (Clarify_tac 1);
 by (rtac Stable_transient_reachable_LeadsTo 1);
 by (res_inst_tac [("k", "k")] transient_lemma 2);
-by (etac Disjoint_States_eq 2);
 by (force_tac (claset() addDs [impOfSubs Increasing_size,
 			       impOfSubs Increasing_Stable_less],
 	       simpset()) 1);
 by (rtac (make_elim (lemma1 RS guaranteesD)) 1);
  by (Blast_tac 1);
-by (etac Disjoint_States_eq 1);
 by (auto_tac (claset(),
 	      simpset() addsimps [Invariant_eq_always, giv_meets_ask_def]));
 by (ALLGOALS Force_tac);
--- a/src/HOL/UNITY/Client.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Client.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -48,8 +48,7 @@
 		          size (ask s) = size (rel s))}"
 
   Cli_prg :: state program
-    "Cli_prg == mk_program (UNIV,
-			    {s. tok s : atMost Max &
+    "Cli_prg == mk_program ({s. tok s : atMost Max &
 			        giv s = [] &
 			        ask s = [] &
 			        rel s = []},
--- a/src/HOL/UNITY/Common.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Common.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -30,13 +30,13 @@
 
 (*** Some programs that implement the safety property above ***)
 
-Goal "SKIP UNIV : constrains {m} (maxfg m)";
+Goal "SKIP : constrains {m} (maxfg m)";
 by (simp_tac (simpset() addsimps [constrains_def, maxfg_def, le_max_iff_disj,
 				  fasc]) 1);
 result();
 
 (*This one is  t := ftime t || t := gtime t*)
-Goal "mk_program (UNIV, UNIV, {range(%t.(t,ftime t)), range(%t.(t,gtime t))}) \
+Goal "mk_program (UNIV, {range(%t.(t,ftime t)), range(%t.(t,gtime t))}) \
 \      : constrains {m} (maxfg m)";
 by (simp_tac (simpset() addsimps [constrains_def, maxfg_def, 
 				  le_max_iff_disj, fasc]) 1);
@@ -44,7 +44,7 @@
 result();
 
 (*This one is  t := max (ftime t) (gtime t)*)
-Goal "mk_program (UNIV, UNIV, {range(%t.(t, max (ftime t) (gtime t)))}) \
+Goal "mk_program (UNIV, {range(%t.(t, max (ftime t) (gtime t)))}) \
 \      : constrains {m} (maxfg m)";
 by (simp_tac 
     (simpset() addsimps [constrains_def, maxfg_def, max_def, gasc]) 1);
@@ -52,8 +52,7 @@
 result();
 
 (*This one is  t := t+1 if t <max (ftime t) (gtime t) *)
-Goal "mk_program (UNIV, UNIV, \
-\                 { {(t, Suc t) | t. t < max (ftime t) (gtime t)} }) \
+Goal "mk_program (UNIV, { {(t, Suc t) | t. t < max (ftime t) (gtime t)} }) \
 \      : constrains {m} (maxfg m)";
 by (simp_tac 
     (simpset() addsimps [constrains_def, maxfg_def, max_def, gasc]) 1);
--- a/src/HOL/UNITY/Comp.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Comp.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -16,33 +16,31 @@
 
 (*** component ***)
 
-Goalw [component_def] "component (SKIP (States F)) F";
+Goalw [component_def] "component SKIP F";
 by (force_tac (claset() addIs [Join_SKIP_left], simpset()) 1);
 qed "component_SKIP";
 
 Goalw [component_def] "component F F";
-by (force_tac (claset() addIs [Join_SKIP_right], simpset()) 1);
+by (blast_tac (claset() addIs [Join_SKIP_right]) 1);
 qed "component_refl";
 
 AddIffs [component_SKIP, component_refl];
 
-Goalw [component_def] "States F = States G ==> component F (F Join G)";
+Goalw [component_def] "component F (F Join G)";
 by (Blast_tac 1);
 qed "component_Join1";
 
-Goalw [component_def] "States F = States G ==> component G (F Join G)";
+Goalw [component_def] "component G (F Join G)";
 by (simp_tac (simpset() addsimps [Join_commute]) 1);
-by (dtac sym 1);
 by (Blast_tac 1);
 qed "component_Join2";
 
-Goalw [component_def, eqStates_def] 
-    "[| i : I;  eqStates I F |] ==> component (F i) (JN i:I. (F i))";
-by (force_tac (claset() addIs [JN_absorb], simpset()) 1);
+Goalw [component_def] "i : I ==> component (F i) (JN i:I. (F i))";
+by (blast_tac (claset() addIs [JN_absorb]) 1);
 qed "component_JN";
 
 Goalw [component_def] "[| component F G; component G H |] ==> component F H";
-by (force_tac (claset() addIs [Join_assoc RS sym], simpset()) 1);
+by (blast_tac (claset() addIs [Join_assoc RS sym]) 1);
 qed "component_trans";
 
 Goalw [component_def] "component F G ==> Acts F <= Acts G";
@@ -53,46 +51,34 @@
 by Auto_tac;
 qed "component_Init";
 
-Goalw [component_def] "component F G ==> States F = States G";
-by Auto_tac;
-qed "component_States";
-
 Goal "[| component F G; component G F |] ==> F=G";
-by (blast_tac (claset() addSIs [program_equalityI, component_States, 
+by (blast_tac (claset() addSIs [program_equalityI, 
 				component_Init, component_Acts]) 1);
 qed "component_anti_sym";
 
 Goalw [component_def]
       "component F H = (EX G. F Join G = H & Disjoint F G)";
-by (blast_tac (claset() addSIs [Disjoint_States_eq, 
-				Diff_Disjoint, Join_Diff2]) 1);
+by (blast_tac (claset() addSIs [Diff_Disjoint, Join_Diff2]) 1);
 qed "component_eq";
 
+
 (*** existential properties ***)
 
 Goalw [ex_prop_def]
-     "[| ex_prop X;  finite GG |] \
-\     ==> eqStates GG (%x. x) --> GG Int X ~= {} --> (JN G:GG. G) : X";
+     "[| ex_prop X; finite GG |] ==> GG Int X ~= {} --> (JN G:GG. G) : X";
 by (etac finite_induct 1);
-by (Simp_tac 1);
-by (rename_tac "GG J" 1);
-by (full_simp_tac (simpset() addsimps [Int_insert_left]) 1);
-by (dres_inst_tac [("x","J")] spec 1);
-by (Force_tac 1);
+by (auto_tac (claset(), simpset() addsimps [Int_insert_left]));
 qed_spec_mp "ex1";
 
 Goalw [ex_prop_def]
-     "ALL GG. finite GG & eqStates GG (%x. x) & GG Int X ~= {} \
-\       --> (JN G:GG. G) : X ==> ex_prop X";
+     "ALL GG. finite GG & GG Int X ~= {} --> (JN G:GG. G) : X ==> ex_prop X";
 by (Clarify_tac 1);
 by (dres_inst_tac [("x", "{F,G}")] spec 1);
 by Auto_tac;
 qed "ex2";
 
 (*Chandy & Sanders take this as a definition*)
-Goal "ex_prop X = \
-\       (ALL GG. finite GG & eqStates GG (%x. x) & GG Int X ~= {} \
-\          --> (JN G:GG. G) : X)";
+Goal "ex_prop X = (ALL GG. finite GG & GG Int X ~= {} --> (JN G:GG. G) : X)";
 by (blast_tac (claset() addIs [ex1,ex2]) 1);
 qed "ex_prop_finite";
 
@@ -103,7 +89,6 @@
 by (Blast_tac 1);
 by Safe_tac;
 by (stac Join_commute 2);
-by (dtac sym 2);
 by (ALLGOALS Blast_tac);
 qed "ex_prop_equiv";
 
@@ -111,15 +96,13 @@
 (*** universal properties ***)
 
 Goalw [uv_prop_def]
-     "[| uv_prop X; finite GG |] \
-\     ==> eqStates GG (%x. x) --> GG <= X --> (JN G:GG. G) : X";
+     "[| uv_prop X; finite GG |] ==> GG <= X --> (JN G:GG. G) : X";
 by (etac finite_induct 1);
 by (auto_tac (claset(), simpset() addsimps [Int_insert_left]));
 qed_spec_mp "uv1";
 
 Goalw [uv_prop_def]
-     "ALL GG. finite GG & eqStates GG (%x. x) & GG <= X \
-\       --> (JN G:GG. G) : X  ==> uv_prop X";
+     "ALL GG. finite GG & GG <= X --> (JN G:GG. G) : X  ==> uv_prop X";
 by (rtac conjI 1);
 by (Clarify_tac 2);
 by (dres_inst_tac [("x", "{F,G}")] spec 2);
@@ -128,8 +111,7 @@
 qed "uv2";
 
 (*Chandy & Sanders take this as a definition*)
-Goal "uv_prop X = (ALL GG. finite GG & eqStates GG (%x. x) & GG <= X \
-\       --> (JN G:GG. G) : X)";
+Goal "uv_prop X = (ALL GG. finite GG & GG <= X --> (JN G:GG. G) : X)";
 by (blast_tac (claset() addIs [uv1,uv2]) 1);
 qed "uv_prop_finite";
 
@@ -199,7 +181,7 @@
 qed "guaranteesI";
 
 Goalw [guarantees_def, component_def]
-     "[| F : X guarantees Y;  F Join G : X;  States F = States G |] \
+     "[| F : X guarantees Y;  F Join G : X;  compatible{F,G} |] \
 \     ==> F Join G : Y";
 by (Blast_tac 1);
 qed "guaranteesD";
@@ -227,42 +209,34 @@
 qed "refines_trans";
 
 Goalw [strict_ex_prop_def]
-     "[| strict_ex_prop X;  States F = States G |] \
-\     ==> (ALL H. States F = States H & F Join H : X --> G Join H : X) = \
-\         (F:X --> G:X)";
-by Safe_tac;
+     "strict_ex_prop X \
+\     ==> (ALL H. F Join H : X --> G Join H : X) = (F:X --> G:X)";
 by (Blast_tac 1);
-by Auto_tac;
 qed "strict_ex_refine_lemma";
 
 Goalw [strict_ex_prop_def]
-     "[| strict_ex_prop X;  States F = States G |] \
-\     ==> (ALL H. States F = States H & F Join H : welldef & F Join H : X \
-\           --> G Join H : X) = \
+     "strict_ex_prop X \
+\     ==> (ALL H. F Join H : welldef & F Join H : X --> G Join H : X) = \
 \         (F: welldef Int X --> G:X)";
 by Safe_tac;
-by (eres_inst_tac [("x","SKIP ?A"), ("P", "%H. ?PP H --> ?RR H")] allE 1);
+by (eres_inst_tac [("x","SKIP"), ("P", "%H. ?PP H --> ?RR H")] allE 1);
 by (auto_tac (claset() addDs [Join_welldef_D1, Join_welldef_D2], simpset()));
 qed "strict_ex_refine_lemma_v";
 
-Goal "[| strict_ex_prop X;  States F = States G;  \
-\        ALL H. States F = States H & F Join H : welldef Int X \
-\          --> G Join H : welldef |] \
+Goal "[| strict_ex_prop X;  \
+\        ALL H. F Join H : welldef Int X --> G Join H : welldef |] \
 \     ==> (G refines F wrt X) = (G iso_refines F wrt X)";
-by (dtac sym 1);
-by (res_inst_tac [("x","SKIP (States G)")] allE 1
+by (res_inst_tac [("x","SKIP")] allE 1
     THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [refines_def, iso_refines_def,
-					   strict_ex_refine_lemma_v]) 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [refines_def, iso_refines_def,
+			 strict_ex_refine_lemma_v]) 1);
 qed "ex_refinement_thm";
 
 
-(***
-
-
 Goalw [strict_uv_prop_def]
      "strict_uv_prop X \
-\     ==> (ALL H. States F = States H & F Join H : X --> G Join H : X) = (F:X --> G:X)";
+\     ==> (ALL H. F Join H : X --> G Join H : X) = (F:X --> G:X)";
 by (Blast_tac 1);
 qed "strict_uv_refine_lemma";
 
@@ -284,4 +258,3 @@
 by (asm_full_simp_tac (simpset() addsimps [refines_def, iso_refines_def,
 					   strict_uv_refine_lemma_v]) 1);
 qed "uv_refinement_thm";
-***)
--- a/src/HOL/UNITY/Comp.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Comp.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -6,11 +6,6 @@
 Composition
 
 From Chandy and Sanders, "Reasoning About Program Composition"
-
-QUESTIONS:
-  refines_def: needs the States F = States G?
-
-  uv_prop, component: should be States F = States (F Join G)
 *)
 
 Comp = Union +
@@ -21,29 +16,23 @@
     case, proving equivalence with Chandy and Sanders's n-ary definitions*)
 
   ex_prop  :: 'a program set => bool
-   "ex_prop X ==
-      ALL F G. (F:X | G: X) & States F = States G --> (F Join G) : X"
+   "ex_prop X == ALL F G. F:X | G: X --> (F Join G) : X"
 
   strict_ex_prop  :: 'a program set => bool
-   "strict_ex_prop X ==
-      ALL F G. States F = States G --> (F:X | G: X) = (F Join G : X)"
+   "strict_ex_prop X == ALL F G. (F:X | G: X) = (F Join G : X)"
 
   uv_prop  :: 'a program set => bool
-   "uv_prop X ==
-      SKIP UNIV : X &
-      (ALL F G. F:X & G: X & States F = States G --> (F Join G) : X)"
+   "uv_prop X == SKIP : X & (ALL F G. F:X & G: X --> (F Join G) : X)"
 
   strict_uv_prop  :: 'a program set => bool
-   "strict_uv_prop X ==
-      SKIP UNIV : X &
-      (ALL F G. States F = States G --> (F:X & G: X) = (F Join G : X))"
+   "strict_uv_prop X == SKIP : X & (ALL F G. (F:X & G: X) = (F Join G : X))"
 
   (*Ill-defined programs can arise through "Join"*)
   welldef :: 'a program set  
    "welldef == {F. Init F ~= {}}"
 
   component :: ['a program, 'a program] => bool
-   "component F H == EX G. F Join G = H & States F = States G"
+   "component F H == EX G. F Join G = H"
 
   guarantees :: ['a program set, 'a program set] => 'a program set (infixl 65)
    "X guarantees Y == {F. ALL H. component F H --> H:X --> H:Y}"
@@ -51,9 +40,7 @@
   refines :: ['a program, 'a program, 'a program set] => bool
 			("(3_ refines _ wrt _)" [10,10,10] 10)
    "G refines F wrt X ==
-      States F = States G &
-      (ALL H. States F = States H & (F Join H) : welldef Int X
-        --> G Join H : welldef Int X)"
+      ALL H. (F Join H) : welldef Int X --> G Join H : welldef Int X"
 
   iso_refines :: ['a program, 'a program, 'a program set] => bool
 			("(3_ iso'_refines _ wrt _)" [10,10,10] 10)
--- a/src/HOL/UNITY/Constrains.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Constrains.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -89,25 +89,24 @@
 by (blast_tac (claset() addIs [constrains_weaken]) 1);
 qed "ball_Constrains_INT";
 
-Goal "[| F : Constrains A A'; A <= reachable F |] ==> A <= A'";
+Goal "F : Constrains A A' ==> reachable F Int A <= A'";
 by (asm_full_simp_tac (simpset() addsimps [Constrains_def]) 1);
 by (dtac constrains_imp_subset 1);
-by (auto_tac (claset() addIs [impOfSubs reachable_subset_States],
-	      simpset() addsimps [Int_subset_iff, Int_lower1]));
+by (ALLGOALS
+    (full_simp_tac (simpset() addsimps [Int_subset_iff, Int_lower1])));
 qed "Constrains_imp_subset";
 
 Goalw [Constrains_def]
     "[| F : Constrains A B; F : Constrains B C |]   \
 \    ==> F : Constrains A C";
-by (blast_tac (claset() addIs [impOfSubs reachable_subset_States,
-			       constrains_trans, constrains_weaken]) 1);
+by (blast_tac (claset() addIs [constrains_trans, constrains_weaken]) 1);
 qed "Constrains_trans";
 
-Goal "[| F : Constrains A (A' Un B); F : Constrains B B' |] \
-\     ==> F : Constrains A (A' Un B')";
-by (asm_full_simp_tac (simpset() addsimps [Constrains_def, Int_Un_distrib]) 1);
-by (blast_tac (claset() addIs [impOfSubs reachable_subset_States,
-			       constrains_cancel]) 1);
+Goalw [Constrains_def, constrains_def]
+   "[| F : Constrains A (A' Un B); F : Constrains B B' |] \
+\   ==> F : Constrains A (A' Un B')";
+by (Clarify_tac 1);
+by (Blast_tac 1);
 qed "Constrains_cancel";
 
 
--- a/src/HOL/UNITY/Handshake.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Handshake.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -11,7 +11,6 @@
 (*split_all_tac causes a big blow-up*)
 claset_ref() := claset() delSWrapper record_split_name;
 
-Addsimps [F_def RS def_prg_States, G_def RS def_prg_States];
 Addsimps [F_def RS def_prg_Init, G_def RS def_prg_Init];
 program_defs_ref := [F_def, G_def];
 
@@ -21,6 +20,7 @@
 Addsimps (thms"state.update_defs");
 Addsimps [simp_of_set invFG_def];
 
+
 Goal "(F Join G) : Invariant invFG";
 by (rtac InvariantI 1);
 by (Force_tac 1);
@@ -32,7 +32,7 @@
 qed "invFG";
 
 Goal "(F Join G) : LeadsTo ({s. NF s = k} - {s. BB s}) \
-\                              ({s. NF s = k} Int {s. BB s})";
+\                          ({s. NF s = k} Int {s. BB s})";
 by (rtac (ensures_stable_Join1 RS leadsTo_Basis RS leadsTo_imp_LeadsTo) 1);
 by (ensures_tac "cmdG" 2);
 by (constrains_tac 1);
--- a/src/HOL/UNITY/Handshake.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Handshake.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -21,14 +21,14 @@
     "cmdF == {(s,s'). s' = s (|NF:= Suc(NF s), BB:=False|) & BB s}"
 
   F :: "state program"
-    "F == mk_program (UNIV, {s. NF s = 0 & BB s}, {cmdF})"
+    "F == mk_program ({s. NF s = 0 & BB s}, {cmdF})"
 
   (*G's program*)
   cmdG :: "(state*state) set"
     "cmdG == {(s,s'). s' = s (|NG:= Suc(NG s), BB:=True|) & ~ BB s}"
 
   G :: "state program"
-    "G == mk_program (UNIV, {s. NG s = 0 & BB s}, {cmdG})"
+    "G == mk_program ({s. NG s = 0 & BB s}, {cmdG})"
 
   (*the joint invariant*)
   invFG :: "state set"
--- a/src/HOL/UNITY/Lift.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Lift.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -116,8 +116,7 @@
 
   Lprg :: state program
     (*for the moment, we OMIT button_press*)
-    "Lprg == mk_program (UNIV,
-			 {s. floor s = Min & ~ up s & move s & stop s &
+    "Lprg == mk_program ({s. floor s = Min & ~ up s & move s & stop s &
 		          ~ open s & req s = {}},
 			 {request_act, open_act, close_act,
 			  req_up, req_down, move_up, move_down})"
@@ -160,5 +159,6 @@
   assumes
     Min_le_n    "Min <= n"
     n_le_Max    "n <= Max"
+  defines
 
 end
--- a/src/HOL/UNITY/Mutex.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Mutex.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -52,8 +52,7 @@
     "cmd4V == {(s,s'). s' = s (|PP:=False, NN:=#0|) & NN s = #4}"
 
   Mprg :: state program
-    "Mprg == mk_program (UNIV,
-			 {s. ~ UU s & ~ VV s & MM s = #0 & NN s = #0},
+    "Mprg == mk_program ({s. ~ UU s & ~ VV s & MM s = #0 & NN s = #0},
 			 {cmd0U, cmd1U, cmd2U, cmd3U, cmd4U, 
 			  cmd0V, cmd1V, cmd2V, cmd3V, cmd4V})"
 
--- a/src/HOL/UNITY/NSP_Bad.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/NSP_Bad.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -54,6 +54,6 @@
 constdefs
   Nprg :: state program
     (*Initial trace is empty*)
-    "Nprg == mk_program(UNIV, {[]}, {Fake, NS1, NS2, NS3})"
+    "Nprg == mk_program({[]}, {Fake, NS1, NS2, NS3})"
 
 end
--- a/src/HOL/UNITY/PPROD.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/PPROD.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -7,241 +7,6 @@
 
 val rinst = read_instantiate_sg (sign_of thy);
 
-    (*** General lemmas ***)
-
-Goalw [sharing_def] "((x,y): sharing Rsh A) = (x: A & y : range (Rsh x))";
-by Auto_tac;
-qed "sharing_iff";
-AddIffs [sharing_iff]; 
-
-Goalw [sharing_def] "(sharing Rsh A <= sharing Rsh B) = (A <= B)";
-by Auto_tac;
-qed "sharing_subset_iff";
-AddIffs [sharing_subset_iff]; 
-
-Goalw [sharing_def] "sharing Rsh (A Un B) = sharing Rsh A Un sharing Rsh B";
-by Auto_tac;
-qed "sharing_Un_distrib";
-
-Goalw [sharing_def] "sharing Rsh (A Int B) = sharing Rsh A Int sharing Rsh B";
-by Auto_tac;
-qed "sharing_Int_distrib";
-
-Goalw [sharing_def] "sharing Rsh (A - B) = sharing Rsh A - sharing Rsh B";
-by Auto_tac;
-qed "sharing_Diff_distrib";
-
-Goalw [sharing_def] "sharing Rsh (Union A) = (UN X:A. sharing Rsh X)";
-by (Blast_tac 1);
-qed "sharing_Union";
-
-Goal "(sharing Rsh A <= - sharing Rsh B) = (A <= - B)";
-by (force_tac (claset(),simpset() addsimps [sharing_def, Image_iff]) 1);
-qed "Lcopy_subset_Compl_eq";
-
-Goal "(((a,b), (a',b')) : Lcopy_act Rsh act) = \
-\     ((a,a'):act & Rsh a b = b & Rsh a' b = b')";
-by (simp_tac (simpset() addsimps [Lcopy_act_def]) 1);
-qed "mem_Lcopy_act_iff";
-AddIffs [mem_Lcopy_act_iff]; 
-
-(*NEEDED????????????????*)
-Goal "[| (a,a'):act; Rsh a b = b |] ==> (((a,b), (a', Rsh a' b)) : Lcopy_act Rsh act)";
-by Auto_tac;
-qed "mem_Lcopy_actI";
-
-
-Goal "act : Acts F \
-\     ==> Lcopy_act Rsh act <= \
-\         sharing Rsh (States F) Times sharing Rsh (States F)";
-by (auto_tac (claset() addIs [range_eqI, sym]
-                       addDs [impOfSubs Acts_subset_Pow_States], 
-              simpset()));
-qed "Lcopy_act_subset_Times";
-
-
-
-Open_locale "Share";
-
-val overwrite = thm "overwrite";
-Addsimps [overwrite];
-
-Goal "Lcopy_act Rsh act ^^ sharing Rsh A = sharing Rsh (act ^^ A)";
-by (force_tac (claset(),simpset() addsimps [sharing_def, Image_iff]) 1);
-qed "Lcopy_act_Image";
-Addsimps [Lcopy_act_Image];
-
-
-
-Goal "(Lcopy_act Rsh act ^^ sharing Rsh A <= sharing Rsh B) = (act ^^ A <= B)";
-by (force_tac (claset(),simpset() addsimps [sharing_def, Image_iff]) 1);
-qed "Lcopy_act_Image_subset_eq";
-
-Goal "Domain (Lcopy_act Rsh act) = sharing Rsh (Domain act)";
-by (auto_tac (claset() addIs [sym], simpset() addsimps [Domain_iff]));
-qed "Domain_Lcopy_act"; 
-
-(*?? needed??
-Goal "(Lcopy_act Rsh act) ^^ (SIGMA x:A. B Int range(Rsh x)) = (SIGMA x: act^^A. Rsh x `` B)";
-by (auto_tac (claset(), simpset() addsimps [Image_iff]));
-qed "Image_Lcopy_act"; 
-**)
-
-Goal "Lcopy_act Rsh (diag A) = diag (sharing Rsh A)";
-by (auto_tac (claset() addIs [sym], simpset()));
-qed "Lcopy_diag";
-
-Addsimps [Lcopy_diag];
-
-
-(**** Lcopy ****)
-
-(*** Basic properties ***)
-
-Goalw [Lcopy_def] "States (Lcopy Rsh F) = sharing Rsh (States F)";
-by Auto_tac;
-qed "States_Lcopy";
-
-Goalw [Lcopy_def] "Init (Lcopy Rsh F) = sharing Rsh (Init F)";
-by (auto_tac (claset() addIs [impOfSubs Init_subset_States], simpset()));
-qed "Init_Lcopy";
-
-Goal "diag (sharing Rsh (States F)) : Lcopy_act Rsh `` Acts F";
-by (rtac image_eqI 1);
-by (rtac diag_in_Acts 2);
-by Auto_tac;
-val lemma = result();
-
-Goal "Acts (Lcopy Rsh F) = (Lcopy_act Rsh `` Acts F)";
-by (auto_tac (claset() addSIs [lemma] 
-                       addDs [impOfSubs Acts_subset_Pow_States], 
-	      simpset() addsimps [Lcopy_def, Lcopy_act_subset_Times, 
-				  image_subset_iff]));
-qed "Acts_Lcopy";
-
-Addsimps [States_Lcopy, Init_Lcopy, Acts_Lcopy];
-
-Goalw [SKIP_def] "Lcopy Rsh (SKIP A) = SKIP(sharing Rsh A)";
-by (rtac program_equalityI 1);
-by Auto_tac;
-qed "Lcopy_SKIP";
-
-
-(*** Safety: constrains, stable ***)
-
-Goal "(Lcopy Rsh F : constrains (sharing Rsh A) (sharing Rsh B)) = \
-\     (F : constrains A B)";
-by (simp_tac (simpset() addsimps [constrains_def, 
-				  Lcopy_act_Image_subset_eq]) 1);
-qed "Lcopy_constrains";
-
-Goal "[| Lcopy Rsh F : constrains A B;  A <= States (Lcopy Rsh F) |]  \
-\     ==> F : constrains (Domain A) (Domain B)";
-by (force_tac (claset() addIs [rev_bexI], 
-	       simpset() addsimps [constrains_def, sharing_def, Image_iff]) 1);
-qed "Lcopy_constrains_Domain";
-
-Goal "(Lcopy Rsh F : stable (sharing Rsh A)) = (F : stable A)";
-by (asm_simp_tac (simpset() addsimps [stable_def, Lcopy_constrains]) 1);
-qed "Lcopy_stable";
-
-Goal "(Lcopy Rsh F : invariant (sharing Rsh A)) = (F : invariant A)";
-by (asm_simp_tac (simpset() addsimps [invariant_def, Lcopy_stable]) 1);
-qed "Lcopy_invariant";
-
-(** Substitution Axiom versions: Constrains, Stable **)
-
-Goal "p : reachable (Lcopy Rsh F) \
-\     ==> (%(a,b). a : reachable F & b : range (Rsh a)) p";
-by (etac reachable.induct 1);
-by (auto_tac
-    (claset() addIs reachable.intrs,
-     simpset() addsimps [Acts_Lcopy]));
-qed "reachable_Lcopy_fst";
-
-Goal "(a,b) : reachable (Lcopy Rsh F) \
-\     ==> a : reachable F & b : range (Rsh a)";
-by (force_tac (claset() addSDs [reachable_Lcopy_fst], simpset()) 1);
-qed "reachable_LcopyD";
-
-Goal "reachable (Lcopy Rsh F) = sharing Rsh (reachable F)";
-by (rtac equalityI 1);
-by (force_tac (claset() addSDs [reachable_LcopyD], simpset()) 1);
-by (Clarify_tac 1);
-by (etac reachable.induct 1);
-by (ALLGOALS (force_tac (claset() addIs reachable.intrs, 
-			 simpset())));
-qed "reachable_Lcopy_eq";
-
-Goal "(Lcopy Rsh F : Constrains (sharing Rsh A) (sharing Rsh B)) =  \
-\     (F : Constrains A B)";
-by (simp_tac
-    (simpset() addsimps [Constrains_def, reachable_Lcopy_eq, 
-			 Lcopy_constrains, sharing_Int_distrib RS sym]) 1);
-qed "Lcopy_Constrains";
-
-Goal "(Lcopy Rsh F : Stable (sharing Rsh A)) = (F : Stable A)";
-by (simp_tac (simpset() addsimps [Stable_def, Lcopy_Constrains]) 1);
-qed "Lcopy_Stable";
-
-
-(*** Progress: transient, ensures ***)
-
-Goal "(Lcopy Rsh F : transient (sharing Rsh A)) = (F : transient A)";
-by (auto_tac (claset(),
-	      simpset() addsimps [transient_def, Lcopy_subset_Compl_eq,
-				  Domain_Lcopy_act]));
-qed "Lcopy_transient";
-
-Goal "(Lcopy Rsh F : ensures (sharing Rsh A) (sharing Rsh B)) = \
-\     (F : ensures A B)";
-by (simp_tac
-    (simpset() addsimps [ensures_def, Lcopy_constrains, Lcopy_transient, 
-			 sharing_Un_distrib RS sym, 
-			 sharing_Diff_distrib RS sym]) 1);
-qed "Lcopy_ensures";
-
-Goal "F : leadsTo A B \
-\     ==> Lcopy Rsh F : leadsTo (sharing Rsh A) (sharing Rsh B)";
-by (etac leadsTo_induct 1);
-by (asm_simp_tac (simpset() addsimps [leadsTo_UN, sharing_Union]) 3);
-by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
-by (asm_simp_tac (simpset() addsimps [leadsTo_Basis, Lcopy_ensures]) 1);
-qed "leadsTo_imp_Lcopy_leadsTo";
-
-(**
-    Goal "Lcopy Rsh F : ensures A B ==> F : ensures (Domain A) (Domain B)";
-    by (full_simp_tac
-	(simpset() addsimps [ensures_def, Lcopy_constrains, Lcopy_transient, 
-			     Domain_Un_eq RS sym,
-			     sharing_Un_distrib RS sym, 
-			     sharing_Diff_distrib RS sym]) 1);
-    by (safe_tac (claset() addSDs [Lcopy_constrains_Domain]));
-    by (etac constrains_weaken_L 1);
-    by (Blast_tac 1);
-    (*NOT able to prove this:
-    Lcopy Rsh F : ensures A B ==> F : ensures (Domain A) (Domain B)
-     1. [| Lcopy Rsh F : transient (A - B);
-	   F : constrains (Domain (A - B)) (Domain (A Un B)) |]
-	==> F : transient (Domain A - Domain B)
-    **)
-
-
-    Goal "Lcopy Rsh F : leadsTo AU BU ==> F : leadsTo (Domain AU) (Domain BU)";
-    by (etac leadsTo_induct 1);
-    by (full_simp_tac (simpset() addsimps [Domain_Union]) 3);
-    by (blast_tac (claset() addIs [leadsTo_UN]) 3);
-    by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
-    by (rtac leadsTo_Basis 1);
-    (*...and so CANNOT PROVE THIS*)
-
-
-    (*This also seems impossible to prove??*)
-    Goal "(Lcopy Rsh F : leadsTo (sharing Rsh A) (sharing Rsh B)) = \
-    \     (F : leadsTo A B)";
-**)
-
-
 (**** PPROD ****)
 
 (*** Basic properties ***)
@@ -251,40 +16,20 @@
 qed "lift_set_iff";
 AddIffs [lift_set_iff];
 
-Goalw [lift_act_def] "lift_act i (diag A) = (diag (lift_set i A))";
+Goalw [lift_act_def] "lift_act i Id = Id";
 by Auto_tac;
-qed "lift_act_diag";
-Addsimps [lift_act_diag];
+qed "lift_act_Id";
+Addsimps [lift_act_Id];
 
-Goalw [lift_prog_def] "States (lift_prog i F) = (lift_set i (States F))";
+Goalw [lift_prog_def] "Init (lift_prog i F) = {f. f i : Init F}";
 by Auto_tac;
-qed "States_lift_prog";
-Addsimps [States_lift_prog];
-
-Goalw [lift_prog_def] "Init (lift_prog i F) = (lift_set i (Init F))";
-by (auto_tac (claset() addIs [impOfSubs Init_subset_States], simpset()));
 qed "Init_lift_prog";
 Addsimps [Init_lift_prog];
 
-Goalw [lift_act_def]
-     "act : Acts F \
-\     ==> lift_act i act <= lift_set i (States F) Times lift_set i (States F)";
-by (force_tac (claset()  addIs [range_eqI, sym]
-                    addDs [impOfSubs Acts_subset_Pow_States], 
-              simpset()) 1);
-qed "lift_act_subset_Times";
-
 Goalw [lift_prog_def] "Acts (lift_prog i F) = lift_act i `` Acts F";
-by (auto_tac (claset() addIs [diag_in_Acts RSN (2,image_eqI)], 
-	      simpset() addsimps [lift_act_subset_Times, 
-				  image_subset_iff]));
+by (auto_tac (claset() addIs [Id_in_Acts RSN (2,image_eqI)], simpset()));
 qed "Acts_lift_prog";
 
-Goalw [PPROD_def] "States (PPROD I F) = (INT i:I. lift_set i (States (F i)))";
-by Auto_tac;
-qed "States_PPROD";
-Addsimps [States_PPROD];
-
 Goalw [PPROD_def] "Init (PPROD I F) = (INT i:I. lift_set i (Init (F i)))";
 by Auto_tac;
 qed "Init_PPROD";
@@ -296,12 +41,7 @@
 qed "lift_act_eq";
 AddIffs [lift_act_eq];
 
-Goalw [eqStates_def] "eqStates I (%i. lift_prog i F)";
-
-Goalw [eqStates_def] "eqStates I (%i. lift_prog i (F i))";
-
-Goal "[| eqStates I (%i. lift_prog i (F i));  i: I |] \
-\     ==> Acts (PPROD I F) = (UN i:I. lift_act i `` Acts (F i))";
+Goal "i: I ==> Acts (PPROD I F) = (UN i:I. lift_act i `` Acts (F i))";
 by (auto_tac (claset(),
 	      simpset() addsimps [PPROD_def, Acts_lift_prog, Acts_JN]));
 qed "Acts_PPROD";
@@ -577,37 +317,6 @@
 
 (*** guarantees properties ***)
 
-(*We only need act2=Id but the condition used is very weak*)
-Goal "(x,y): act2 ==> fst_act (act1 Lcopy_act act2) = act1";
-by (auto_tac (claset(), simpset() addsimps [fst_act_def, Lcopy_act_def]));
-qed "fst_act_Lcopy_act";
-Addsimps [fst_act_Lcopy_act];
-
-
-Goal "(Lcopy Rsh F) Join G = Lcopy H ==> EX J. H = F Join J";
-by (etac program_equalityE 1);
-by (auto_tac (claset(), simpset() addsimps [Acts_Join]));
-by (res_inst_tac 
-     [("x", "mk_program(Domain (Init G), fst_act `` Acts G)")] exI 1);
-by (rtac program_equalityI 1);
-(*Init*)
-by (simp_tac (simpset() addsimps [Acts_Join]) 1);
-by (blast_tac (claset() addEs [equalityE]) 1);
-(*Now for the Actions*)
-by (dres_inst_tac [("f", "op `` fst_act")] arg_cong 1);
-by (asm_full_simp_tac 
-    (simpset() addsimps [insert_absorb, Acts_Lcopy, Acts_Join,
-			 image_Un, image_compose RS sym, o_def]) 1);
-qed "Lcopy_Join_eq_Lcopy_D";
-
-
-Goal "F : X guarantees Y \
-\     ==> Lcopy Rsh F : (Lcopy `` X) guarantees (Lcopy `` Y)";
-by (rtac guaranteesI 1);
-by Auto_tac;
-by (blast_tac (claset() addDs [Lcopy_Join_eq_Lcopy_D, guaranteesD]) 1);
-qed "Lcopy_guarantees";
-
 
 Goal "drop_act i (lift_act i act) = act";
 by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] exI],
--- a/src/HOL/UNITY/PPROD.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/PPROD.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -4,35 +4,14 @@
     Copyright   1998  University of Cambridge
 
 General products of programs (Pi operation), for replicating components.
-Also merging of state sets.
-
-The idea of Rsh is to represent sharing in the Right part.
-If x and y are states then (Rsh x y) updates y to agree with variables shared
-with x.  Therefore Rsh x (Rsh x y) = Rsh x y.  The pair (x,y)
-is a valid state of the composite program if and only if y = Rsh x y.
-
-Needs Rcopy; try to do by swapping (symmetry argument)
-  instead of repeating all Lcopy proofs.
 *)
 
 PPROD = Union + Comp +
 
 constdefs
 
-  sharing :: "[['a,'b]=>'b, 'a set] => ('a*'b) set"
-    "sharing Rsh A == SIGMA x: A. range (Rsh x)"
-
-  Lcopy_act :: "[['a,'b]=>'b, ('a*'a) set] => (('a*'b) * ('a*'b)) set"
-    "Lcopy_act Rsh act == {((x,y),(x',y')). (x,x'): act & Rsh x y = y &
-			                    Rsh x' y = y'}"
-
-  fst_act :: "(('a*'b) * ('a*'b)) set => ('a*'a) set"
-    "fst_act act == (%((x,y),(x',y')). (x,x')) `` act"
-
-  Lcopy :: "[['a,'b]=>'b, 'a program] => ('a*'b) program"
-    "Lcopy Rsh F == mk_program (sharing Rsh (States F),
-			        sharing Rsh (Init F),
-			        Lcopy_act Rsh `` Acts F)"
+  (**possible to force all acts to be included in common initial state
+      (by intersection) ??*)
 
   lift_act :: "['a, ('b*'b) set] => (('a=>'b) * ('a=>'b)) set"
     "lift_act i act == {(f,f'). EX s'. f' = f(i:=s') & (f i, s') : act}"
@@ -45,8 +24,7 @@
 
   lift_prog :: "['a, 'b program] => ('a => 'b) program"
     "lift_prog i F ==
-       mk_program (lift_set i (States F),
-		   lift_set i (Init F),
+       mk_program (lift_set i (Init F),
 		   lift_act i `` Acts F)"
 
   (*products of programs*)
@@ -59,12 +37,4 @@
 translations
   "PPI x:A. B"   == "PPROD A (%x. B)"
 
-
-locale Share =
-  fixes 
-    Rsh	:: ['a,'b]=>'b
-  assumes
-    (*the last update (from the other side) takes precedence*)
-    overwrite "Rsh x (Rsh x' y) = Rsh x y"
-
 end
--- a/src/HOL/UNITY/ROOT.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/ROOT.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -29,7 +29,8 @@
 time_use_thy "Comp";
 time_use_thy "Client";
 (**
-time_use_thy "PPX";
+time_use_thy "Extend";
+time_use_thy "PPROD";
 **)
 
 add_path "../Auth";	(*to find Public.thy*)
--- a/src/HOL/UNITY/Reach.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Reach.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -24,7 +24,7 @@
     "asgt u v == {(s,s'). s' = s(v:= s u | s v)}"
 
   Rprg :: state program
-    "Rprg == mk_program (UNIV, {%v. v=init}, UN (u,v): edges. {asgt u v})"
+    "Rprg == mk_program ({%v. v=init}, UN (u,v): edges. {asgt u v})"
 
   reach_invariant :: state set
     "reach_invariant == {s. (ALL v. s v --> (init, v) : edges^*) & s init}"
--- a/src/HOL/UNITY/SubstAx.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/SubstAx.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -79,7 +79,7 @@
 
 val prems = 
 Goal "(!!i. i : I ==> F : LeadsTo (A i) B) ==> F : LeadsTo (UN i:I. A i) B";
-by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
+by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1);
 by (blast_tac (claset() addIs (LeadsTo_Union::prems)) 1);
 qed "LeadsTo_UN";
 
@@ -190,7 +190,7 @@
 val prems = 
 Goal "(!! i. i:I ==> F : LeadsTo (A i) (A' i)) \
 \     ==> F : LeadsTo (UN i:I. A i) (UN i:I. A' i)";
-by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
+by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1);
 by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_R] 
                         addIs prems) 1);
 qed "LeadsTo_UN_UN";
--- a/src/HOL/UNITY/Traces.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Traces.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -10,110 +10,41 @@
 *)
 
 
+
 (*** The abstract type of programs ***)
 
-Goalw [restrict_rel_def] "restrict_rel A Id = diag A";
-by (Blast_tac 1);
-qed "restrict_rel_Id";
-Addsimps [restrict_rel_Id];
-
-Goalw [restrict_rel_def] "restrict_rel A (diag B) = diag (A Int B)";
-by (Blast_tac 1);
-qed "restrict_rel_diag";
-Addsimps [restrict_rel_diag];
-
-Goalw [restrict_rel_def]
-    "restrict_rel A (restrict_rel B r) = restrict_rel (A Int B) r";
-by (Blast_tac 1);
-qed "restrict_rel_restrict_rel";
-Addsimps [restrict_rel_restrict_rel];
-
-Goalw [restrict_rel_def] "restrict_rel A r <= A Times A";
-by (Blast_tac 1);
-qed "restrict_rel_subset";
-Addsimps [restrict_rel_subset];
-
-Goalw [restrict_rel_def]
-    "((x,y) : restrict_rel A r) = ((x,y): r & x: A & y: A)";
-by (Blast_tac 1);
-qed "restrict_rel_iff";
-Addsimps [restrict_rel_iff];
-
-Goalw [restrict_rel_def] "r <= A Times A ==> restrict_rel A r = r";
-by (Blast_tac 1);
-qed "restrict_rel_eq";
-Addsimps [restrict_rel_eq];
-
-Goal "acts <= Pow (A Times A) ==> restrict_rel A `` acts = acts";
-by Auto_tac;
-by (rtac image_eqI 1);
-by (assume_tac 2);
-by (set_mp_tac 1);
-by (Force_tac 1);
-qed "restrict_rel_image";
-
-Goalw [restrict_rel_def] "Domain (restrict_rel A r) <= A Int Domain r";
-by (Blast_tac 1);
-qed "Domain_restrict_rel";
-
-Goalw [restrict_rel_def] "restrict_rel A r ^^ B <= A Int (r ^^ B)";
-by (Blast_tac 1);
-qed "Image_restrict_rel";
-
-Addsimps [diag_iff];
-
 val rep_ss = simpset() addsimps 
-                [Int_lower1, image_subset_iff, diag_subset_Times,
-		 States_def, Init_def, Acts_def, 
-		 mk_program_def, Program_def, Rep_Program, 
+                [Init_def, Acts_def, mk_program_def, Program_def, Rep_Program, 
 		 Rep_Program_inverse, Abs_Program_inverse];
 
 
-(** Basic laws guaranteed by the abstract type "program" **)
-
-Goal "Init F <= States F";
+Goal "Id : Acts F";
 by (cut_inst_tac [("x", "F")] Rep_Program 1);
 by (auto_tac (claset(), rep_ss));
-qed "Init_subset_States";
-
-Goal "Acts F <= Pow(States F Times States F)";
-by (cut_inst_tac [("x", "F")] Rep_Program 1);
-by (force_tac (claset(),rep_ss) 1);
-qed "Acts_subset_Pow_States";
+qed "Id_in_Acts";
+AddIffs [Id_in_Acts];
 
-Goal "diag (States F) : Acts F";
-by (cut_inst_tac [("x", "F")] Rep_Program 1);
-by (auto_tac (claset(), rep_ss));
-qed "diag_in_Acts";
-AddIffs [diag_in_Acts];
-
+Goal "insert Id (Acts F) = Acts F";
+by (simp_tac (simpset() addsimps [insert_absorb, Id_in_Acts]) 1);
+qed "insert_Id_Acts";
+AddIffs [insert_Id_Acts];
 
 (** Inspectors for type "program" **)
 
-Goal "States (mk_program (states,init,acts)) = states";
-by (auto_tac (claset(), rep_ss));
-qed "States_eq";
-
-Goal "Init (mk_program (states,init,acts)) = states Int init";
+Goal "Init (mk_program (init,acts)) = init";
 by (auto_tac (claset(), rep_ss));
 qed "Init_eq";
 
-Goal "Acts (mk_program (states,init,acts)) = \
-\     insert (diag states) (restrict_rel states `` acts)";
+Goal "Acts (mk_program (init,acts)) = insert Id acts";
 by (auto_tac (claset(), rep_ss));
-qed "Acts_eq_raw";
-
-Goal "acts <= Pow(states Times states) \
-\     ==> Acts (mk_program (states,init,acts)) = insert (diag states) acts";
-by (asm_simp_tac (simpset() addsimps [Acts_eq_raw, restrict_rel_image]) 1);
 qed "Acts_eq";
 
-Addsimps [States_eq, Acts_eq, Init_eq];
+Addsimps [Acts_eq, Init_eq];
 
 
 (** The notation of equality for type "program" **)
 
-Goal "[| States F = States G; Init F = Init G; Acts F = Acts G |] ==> F = G";
+Goal "[| Init F = Init G; Acts F = Acts G |] ==> F = G";
 by (subgoals_tac ["EX x. Rep_Program F = x",
 		  "EX x. Rep_Program G = x"] 1);
 by (REPEAT (Blast_tac 2));
@@ -124,38 +55,26 @@
 qed "program_equalityI";
 
 val [major,minor] =
-Goal "[| F = G; \
-\        [| States F = States G; Init F = Init G; Acts F = Acts G |] ==> P \
-\     |] ==> P";
+Goal "[| F = G; [| Init F = Init G; Acts F = Acts G |] ==> P |] ==> P";
 by (rtac minor 1);
 by (auto_tac (claset(), simpset() addsimps [major]));
 qed "program_equalityE";
 
+
 (*** These rules allow "lazy" definition expansion 
      They avoid expanding the full program, which is a large expression
 ***)
 
-Goal "[| F == mk_program (states,init,acts) |] \
-\     ==> States F = states";
-by Auto_tac;
-qed "def_prg_States";
-
-Goal "[| F == mk_program (states,init,acts); init <= states |] \
-\     ==> Init F = init";
+Goal "F == mk_program (init,acts) ==> Init F = init";
 by Auto_tac;
 qed "def_prg_Init";
 
-Goal "[| F == mk_program (states,init,acts); \
-\        acts <= Pow(states Times states) |] \
-\     ==> Acts F = insert (diag states) acts";
-by (asm_simp_tac (simpset() addsimps [restrict_rel_image]) 1);
-qed "def_prg_Acts";
-
 (*The program is not expanded, but its Init and Acts are*)
-Goal "[| F == mk_program (states,init,acts); \
-\        init <= states;  acts <= Pow(states Times states) |] \
-\     ==> Init F = init & Acts F = insert (diag states) acts";
-by (asm_simp_tac (HOL_ss addsimps [def_prg_Init, def_prg_Acts]) 1);
+val [rew] = goal thy
+    "[| F == mk_program (init,acts) |] \
+\    ==> Init F = init & Acts F = insert Id acts";
+by (rewtac rew);
+by Auto_tac;
 qed "def_prg_simps";
 
 (*An action is expanded only if a pair of states is being tested against it*)
@@ -175,13 +94,6 @@
 
 (*** traces and reachable ***)
 
-Goal "reachable F <= States F";
-by Safe_tac;
-by (etac reachable.induct 1);
-by (blast_tac (claset() addDs [impOfSubs Acts_subset_Pow_States]) 2);
-by (blast_tac (claset() addIs [impOfSubs Init_subset_States]) 1);
-qed "reachable_subset_States";
-
 Goal "reachable F = {s. EX evs. (s,evs): traces (Init F) (Acts F)}";
 by Safe_tac;
 by (etac traces.induct 2);
--- a/src/HOL/UNITY/Traces.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Traces.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -25,32 +25,17 @@
 
 
 typedef (Program)
-  'a program = "{(states :: 'a set,
-		  init   :: 'a set,
-		  acts   :: ('a * 'a)set set).
-		 init <= states &
-		 acts <= Pow(states Times states) &
-		 diag states : acts}"
+  'a program = "{(init:: 'a set, acts :: ('a * 'a)set set). Id:acts}"
 
 constdefs
-  restrict_rel :: "['a set, ('a * 'a) set] => ('a * 'a) set"
-    "restrict_rel A r == (A Times A) Int r"
-
-  mk_program :: "('a set * 'a set * ('a * 'a)set set) => 'a program"
-    "mk_program ==
-       %(states, init, acts).
-	Abs_Program (states,
-		     states Int init,
-		     restrict_rel states `` (insert Id acts))"
-
-  States :: "'a program => 'a set"
-    "States F == (%(states, init, acts). states) (Rep_Program F)"
+    mk_program :: "('a set * ('a * 'a)set set) => 'a program"
+    "mk_program == %(init, acts). Abs_Program (init, insert Id acts)"
 
   Init :: "'a program => 'a set"
-    "Init F == (%(states, init, acts). init) (Rep_Program F)"
+    "Init F == (%(init, acts). init) (Rep_Program F)"
 
   Acts :: "'a program => ('a * 'a)set set"
-    "Acts F == (%(states, init, acts). acts) (Rep_Program F)"
+    "Acts F == (%(init, acts). acts) (Rep_Program F)"
 
 
 consts reachable :: "'a program => 'a set"
--- a/src/HOL/UNITY/UNITY.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/UNITY.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -14,11 +14,6 @@
 
 (*** General lemmas ***)
 
-    Goal "Pow UNIV = UNIV";
-    by (Blast_tac 1);
-    qed "Pow_UNIV";
-    Addsimps [Pow_UNIV];
-
 Goal "UNIV Times UNIV = UNIV";
 by Auto_tac;
 qed "UNIV_Times_UNIV"; 
@@ -106,23 +101,23 @@
 by (Blast_tac 1);
 qed "ball_constrains_INT";
 
-Goalw [constrains_def] "[| F : constrains A A'; A <= States F |] ==> A <= A'";
-by (Force_tac 1);
+Goalw [constrains_def] "F : constrains A A' ==> A <= A'";
+by Auto_tac;
 qed "constrains_imp_subset";
 
 (*The reasoning is by subsets since "constrains" refers to single actions
   only.  So this rule isn't that useful.*)
-Goal "[| F : constrains A B; F : constrains B C; B <= States F |]   \
+Goalw [constrains_def]
+    "[| F : constrains A B; F : constrains B C |]   \
 \    ==> F : constrains A C";
-by (etac constrains_weaken_R 1);
-by (etac constrains_imp_subset 1);
-by (assume_tac 1);
+by (Blast_tac 1);
 qed "constrains_trans";
 
-Goal "[| F : constrains A (A' Un B); F : constrains B B'; B <= States F|] \
-\     ==> F : constrains A (A' Un B')";
-by (etac constrains_weaken_R 1);
-by (blast_tac (claset() addDs [impOfSubs constrains_imp_subset]) 1);
+Goalw [constrains_def]
+   "[| F : constrains A (A' Un B); F : constrains B B' |] \
+\   ==> F : constrains A (A' Un B')";
+by (Clarify_tac 1);
+by (Blast_tac 1);
 qed "constrains_cancel";
 
 
--- a/src/HOL/UNITY/Union.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Union.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -8,6 +8,7 @@
 From Misra's Chapter 5: Asynchronous Compositions of Programs
 *)
 
+
 Goal "k:I ==> (INT i:I. A i) Int A k = (INT i:I. A i)";
 by (Blast_tac 1);
 qed "INT_absorb2";
@@ -17,23 +18,20 @@
 
 Addcongs [UN_cong, INT_cong];
 
+
 (** SKIP **)
 
-Goal "States (SKIP A) = A";
-by (simp_tac (simpset() addsimps [SKIP_def]) 1);
-qed "States_SKIP";
-
-Goal "Init (SKIP A) = A";
+Goal "Init SKIP = UNIV";
 by (simp_tac (simpset() addsimps [SKIP_def]) 1);
 qed "Init_SKIP";
 
-Goal "Acts (SKIP A) = {diag A}";
+Goal "Acts SKIP = {Id}";
 by (simp_tac (simpset() addsimps [SKIP_def]) 1);
 qed "Acts_SKIP";
 
-Addsimps [States_SKIP, Init_SKIP, Acts_SKIP];
+Addsimps [Init_SKIP, Acts_SKIP];
 
-Goal "reachable (SKIP A) = A";
+Goal "reachable SKIP = UNIV";
 by (force_tac (claset() addEs [reachable.induct]
 			addIs reachable.intrs, simpset()) 1);
 qed "reachable_SKIP";
@@ -43,90 +41,47 @@
 
 (** Join **)
 
-Goal "States (F Join G) = States F Int States G";
+Goal "Init (F Join G) = Init F Int Init G";
 by (simp_tac (simpset() addsimps [Join_def]) 1);
-qed "States_Join";
-
-Goal "Init (F Join G) = Init F Int Init G";
-by (auto_tac (claset() addIs [impOfSubs Init_subset_States],
-	      simpset() addsimps [Join_def]));
 qed "Init_Join";
 
-Goal "States F = States G ==> Acts (F Join G) = Acts F Un Acts G";
-by (subgoal_tac "Acts F Un Acts G <= Pow (States G Times States G)" 1);
-by (blast_tac (claset() addEs [equalityE] 
-                        addDs [impOfSubs Acts_subset_Pow_States]) 2);
+Goal "Acts (F Join G) = Acts F Un Acts G";
 by (auto_tac (claset(), simpset() addsimps [Join_def]));
 qed "Acts_Join";
 
+Addsimps [Init_Join];
+
 
 (** JN **)
 
-Goalw [JOIN_def, SKIP_def] "(JN i:{}. F i) = SKIP UNIV";
+Goalw [JOIN_def, SKIP_def] "(JN i:{}. F i) = SKIP";
 by Auto_tac;
 qed "JN_empty";
 Addsimps [JN_empty];
 
-Goal "States (JN i:I. F i) = (INT i:I. States (F i))";
-by (simp_tac (simpset() addsimps [JOIN_def]) 1);
-qed "States_JN";
+Goal "(JN x:insert a A. B x) = (B a) Join (JN x:A. B x)";
+by (rtac program_equalityI 1);
+by (ALLGOALS (simp_tac (simpset() addsimps [JOIN_def, Join_def])));
+qed "JN_insert";
+Addsimps[JN_empty, JN_insert];
 
 Goal "Init (JN i:I. F i) = (INT i:I. Init (F i))";
-by (auto_tac (claset() addIs [impOfSubs Init_subset_States],
-	      simpset() addsimps [JOIN_def]));
+by (simp_tac (simpset() addsimps [JOIN_def]) 1);
 qed "Init_JN";
 
-(*If I={} then the LHS is (SKIP UNIV) while the RHS is {}.*)
-Goalw [eqStates_def]
-     "[| eqStates I F; i: I |] ==> Acts (JN i:I. F i) = (UN i:I. Acts (F i))";
-by (Clarify_tac 1);
-by (subgoal_tac "(UN i:I. Acts (F i)) <= Pow (St Times St)" 1);
-by (blast_tac (claset() addEs [equalityE] 
-                        addDs [impOfSubs Acts_subset_Pow_States]) 2);
+Goal "Acts (JN i:I. F i) = insert Id (UN i:I. Acts (F i))";
 by (auto_tac (claset(), simpset() addsimps [JOIN_def]));
 qed "Acts_JN";
 
-Goal "eqStates I F \
-\     ==> Acts (JN i:I. F i) = \
-\         (if I={} then {diag UNIV} else (UN i:I. Acts (F i)))";
-by (force_tac (claset(), simpset() addsimps [Acts_JN]) 1);
-qed "Acts_JN_if";
-
-Addsimps [States_Join, Init_Join, States_JN, Init_JN];
-
-
-Goal "(JN x:insert j I. F x) = (F j) Join (JN x:I. F x)";
-by (rtac program_equalityI 1);
-by (ALLGOALS Asm_simp_tac);
-by (asm_simp_tac 
-    (simpset() addsimps [JOIN_def, Join_def, Acts_eq_raw,
-			 image_UNION, image_Un, image_image, Int_assoc]) 1);
-qed "JN_insert";
-Addsimps[JN_insert];
+Addsimps [Init_JN];
 
-Goal "k:I ==> A k Join (JN i:I. A i) = (JN i:I. A i)";
-by (stac (JN_insert RS sym) 1);
-by (asm_simp_tac (simpset() addsimps [insert_absorb]) 1);
-qed "JN_absorb";
-
-
-Goalw [eqStates_def] "eqStates {} F";
-by (Simp_tac 1);
-qed "eqStates_empty";
+val prems = Goalw [JOIN_def]
+    "[| A=B;  !!x. x:B ==> F(x) = G(x) |] ==> \
+\    (JN x:A. F(x)) = (JN x:B. G(x))";
+by (asm_simp_tac (simpset() addsimps prems) 1);
+qed "JN_cong";
 
-Goalw [eqStates_def]
-    "[| eqStates I F; States (F i) = (INT i: I. States (F i)) |] \
-\    ==> eqStates (insert i I) F";
-by Auto_tac;
-qed "eqStates_insertI";
-
-Goalw [eqStates_def]
-    "eqStates (insert i I) F = \
-\    (eqStates I F & (I={} | States (F i) = (INT i: I. States (F i))))";
-by Auto_tac;
-qed "eqStates_insert_eq";
-
-Addsimps [eqStates_empty, eqStates_insert_eq];
+Addcongs [JN_cong];
 
 
 (** Algebraic laws **)
@@ -136,97 +91,86 @@
 qed "Join_commute";
 
 Goal "(F Join G) Join H = F Join (G Join H)";
-by (rtac program_equalityI 1);
-by (ALLGOALS (asm_simp_tac (simpset() addsimps Un_ac@[Int_assoc])));
-by (asm_simp_tac 
-    (simpset() addsimps Un_ac@Int_ac@[Join_def, Acts_eq_raw,
-				      image_Un, image_image]) 1);
-by (Blast_tac 1);
+by (simp_tac (simpset() addsimps Un_ac@[Join_def, Int_assoc]) 1);
 qed "Join_assoc";
  
-Goalw [Join_def, SKIP_def] "States F <= A ==> (SKIP A) Join F = F";
+Goalw [Join_def, SKIP_def] "SKIP Join F = F";
 by (rtac program_equalityI 1);
-by (ALLGOALS
-    (asm_simp_tac (simpset() addsimps 
-		   Int_ac@[Acts_subset_Pow_States RS restrict_rel_image, 
-			   Acts_eq_raw, insert_absorb, Int_absorb1])));
-by (blast_tac (claset() addIs [impOfSubs Init_subset_States]) 1);
+by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb])));
 qed "Join_SKIP_left";
 
-Goal "States F <= A ==> F Join (SKIP A) = F";
-by (stac Join_commute 1);
-by (asm_simp_tac (simpset() addsimps [Join_SKIP_left]) 1);
+Goalw [Join_def, SKIP_def] "F Join SKIP = F";
+by (rtac program_equalityI 1);
+by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb])));
 qed "Join_SKIP_right";
 
 Addsimps [Join_SKIP_left, Join_SKIP_right];
 
 Goalw [Join_def] "F Join F = F";
 by (rtac program_equalityI 1);
-by (ALLGOALS
-    (asm_simp_tac (simpset() addsimps 
-		   [insert_absorb, Acts_subset_Pow_States RS Acts_eq])));
-by (blast_tac (claset() addIs [impOfSubs Init_subset_States]) 1);
+by Auto_tac;
 qed "Join_absorb";
 
 Addsimps [Join_absorb];
 
+
+
+(*** JN laws ***)
+
+
+
+Goal "i : I ==> (JN i:I. A i Join B) = ((JN i:I. A i) Join B)";
+by (auto_tac (claset() addSIs [program_equalityI],
+	      simpset() addsimps [Acts_JN, Acts_Join]));
+qed "JN_Join_miniscope";
+
+Goal "k:I ==> A k Join (JN i:I. A i) = (JN i:I. A i)";
+by (auto_tac (claset() addSIs [program_equalityI],
+	      simpset() addsimps [Acts_JN, Acts_Join]));
+qed "JN_absorb";
+
 Goal "(JN i: I Un J. A i) = ((JN i: I. A i) Join (JN i:J. A i))";
-by (rtac program_equalityI 1);
-by (ALLGOALS (asm_simp_tac (simpset() addsimps [INT_Un])));
-by (asm_simp_tac 
-    (simpset() addsimps Int_ac@
-                        [JOIN_def, Join_def, Acts_eq_raw, UN_Un, INT_Un, 
-			 image_UNION, image_Un, image_image]) 1);
+by (auto_tac (claset() addSIs [program_equalityI],
+	      simpset() addsimps [Acts_JN, Acts_Join]));
 qed "JN_Join";
 
 Goal "i: I ==> (JN i:I. c) = c";
 by (auto_tac (claset() addSIs [program_equalityI],
-	      simpset() addsimps [Acts_JN, eqStates_def]));
+	      simpset() addsimps [Acts_JN]));
 qed "JN_constant";
 
 Goal "(JN i:I. A i Join B i) = (JN i:I. A i)  Join  (JN i:I. B i)";
-by (rtac program_equalityI 1);
-by (ALLGOALS (asm_simp_tac (simpset() addsimps [INT_Int_distrib])));
-by (asm_simp_tac 
-    (simpset() addsimps [JOIN_def, Join_def, Acts_eq_raw, image_UNION, 
-			 rinst [("A","%x. States (A x) Int States (B x)")] 
-			     INT_absorb2, 
-			 image_Un, image_image]) 1);
-by (asm_simp_tac (simpset() addsimps [INT_Int_distrib RS sym] @ Int_ac) 1);
-by (Blast_tac 1);
+by (auto_tac (claset() addSIs [program_equalityI],
+	      simpset() addsimps [Acts_JN, Acts_Join]));
 qed "JN_Join_distrib";
 
-Goal "i : I ==> (JN i:I. A i Join B) = ((JN i:I. A i) Join B)";
-by (asm_simp_tac (simpset() addsimps [JN_Join_distrib, JN_constant]) 1);
-qed "JN_Join_miniscope";
-
 
 (*** Safety: constrains, stable, FP ***)
 
-Goal "[| F : constrains A B;  G : constrains A B |] \
-\     ==> F Join G : constrains A B";
-by (auto_tac (claset(),
-	      simpset() addsimps [constrains_def, Join_def, Acts_eq_raw, 
-				  image_Un]));
-by (REPEAT (force_tac (claset() addSEs [ballE], simpset()) 1));
-qed "constrains_imp_Join_constrains";
+Goalw [constrains_def, JOIN_def]
+    "i : I ==> \
+\    (JN i:I. F i) : constrains A B = (ALL i:I. F i : constrains A B)";
+by (Simp_tac 1);
+by (Blast_tac 1);
+qed "constrains_JN";
 
-Goal "States F = States G ==> \
-\     F Join G : constrains A B = (F : constrains A B & G : constrains A B)";
-by (auto_tac (claset(), simpset() addsimps [constrains_def, Acts_Join]));
+Goal "F Join G : constrains A B =  \
+\     (F : constrains A B & G : constrains A B)";
+by (auto_tac
+    (claset(),
+     simpset() addsimps [constrains_def, Join_def]));
 qed "constrains_Join";
 
 Goal "[| i : I;  ALL i:I. F i : constrains A B |] \
 \     ==> (JN i:I. F i) : constrains A B";
-by (full_simp_tac (simpset() addsimps [constrains_def, JOIN_def, Acts_eq_raw, 
-				  image_Un]) 1);
-by Safe_tac;
-by (REPEAT (force_tac (claset() addSEs [ballE], simpset()) 1));
+by (full_simp_tac (simpset() addsimps [constrains_def, Acts_JN]) 1);
+by (Blast_tac 1);
 qed "constrains_imp_JN_constrains";
 
-Goal "[| i : I;  eqStates I F |] \
+Goal "[| i : I;  compatible (F``I) |] \
 \     ==> (JN i:I. F i) : constrains A B = (ALL i:I. F i : constrains A B)";
-by (auto_tac (claset(), simpset() addsimps [constrains_def, Acts_JN]));
+by (asm_simp_tac (simpset() addsimps [constrains_def, Acts_JN]) 1);
+by (Blast_tac 1);
 qed "constrains_JN";
 
     (**FAILS, I think; see 5.4.1, Substitution Axiom.
@@ -246,114 +190,89 @@
     qed "Constrains_Join";
     **)
 
-Goal "[| F : constrains A A';  G : constrains B B';  States F = States G |] \
+
+Goal "[| F : constrains A A';  G : constrains B B' |] \
 \     ==> F Join G : constrains (A Int B) (A' Un B')";
-by (asm_simp_tac (simpset() addsimps [constrains_Join]) 1);
+by (simp_tac (simpset() addsimps [constrains_Join]) 1);
 by (blast_tac (claset() addIs [constrains_weaken]) 1);
 qed "constrains_Join_weaken";
 
-Goal "[| i : I;  eqStates I F |]  \
-\     ==> (JN i:I. F i) : stable A = (ALL i:I. F i : stable A)";
+Goal "i : I ==> \
+\     (JN i:I. F i) : stable A = (ALL i:I. F i : stable A)";
 by (asm_simp_tac (simpset() addsimps [stable_def, constrains_JN]) 1);
 qed "stable_JN";
 
-Goal "[| ALL i:I. F i : invariant A;  i : I;  eqStates I F |]  \
+Goal "[| ALL i:I. F i : invariant A;  i : I |]  \
 \      ==> (JN i:I. F i) : invariant A";
 by (asm_full_simp_tac (simpset() addsimps [invariant_def, stable_JN]) 1);
 by (Blast_tac 1);
 bind_thm ("invariant_JN_I", ballI RS result());
 
-Goal "States F = States G \
-\     ==> F Join G : stable A = (F : stable A & G : stable A)";
-by (asm_simp_tac (simpset() addsimps [stable_def, constrains_Join]) 1);
+Goal "F Join G : stable A = \
+\     (F : stable A & G : stable A)";
+by (simp_tac (simpset() addsimps [stable_def, constrains_Join]) 1);
 qed "stable_Join";
 
-Goal "[| F : invariant A; G : invariant A; States F = States G |]  \
+Goal "[| F : invariant A; G : invariant A |]  \
 \     ==> F Join G : invariant A";
-by (asm_full_simp_tac (simpset() addsimps [invariant_def, stable_Join]) 1);
+by (full_simp_tac (simpset() addsimps [invariant_def, stable_Join]) 1);
 by (Blast_tac 1);
 qed "invariant_JoinI";
 
-Goal "[| i : I;  eqStates I F |]  \
-\      ==> FP (JN i:I. F i) = (INT i:I. FP (F i))";
+Goal "i : I ==> FP (JN i:I. F i) = (INT i:I. FP (F i))";
 by (asm_simp_tac (simpset() addsimps [FP_def, stable_JN, INTER_def]) 1);
 qed "FP_JN";
 
 
 (*** Progress: transient, ensures ***)
 
-
-Goal "[| (JN i:I. F i) : transient A;  i: I |] ==> EX i:I. F i : transient A";
-by (full_simp_tac (simpset() addsimps [transient_def, JOIN_def, Acts_eq_raw, 
-				  Int_absorb1, restrict_rel_def]) 1);
-by Safe_tac;
-by (REPEAT (blast_tac (claset() addSIs [rev_bexI]) 1));
-qed "transient_JN_imp_transient";
-
-Goal "[| i : I;  eqStates I F |]  \
-\     ==> (JN i:I. F i) : transient A = (EX i:I. F i : transient A)";
-by (auto_tac (claset() addSIs [transient_JN_imp_transient], simpset()));
-by (auto_tac (claset(), simpset() addsimps [transient_def, Acts_JN]));
+Goal "i : I ==> \
+\   (JN i:I. F i) : transient A = (EX i:I. F i : transient A)";
+by (auto_tac (claset(),
+	      simpset() addsimps [transient_def, JOIN_def]));
 qed "transient_JN";
 
-Goal "F Join G : transient A ==> \
-\     F : transient A | G : transient A";
-by (full_simp_tac (simpset() addsimps [transient_def, Join_def, Acts_eq_raw, 
-				       restrict_rel_def]) 1);
-by Safe_tac;
-(*delete act:Acts F possibility*)
-by (rtac FalseE 3);  
-(*delete act:Acts G possibility*)
-by (thin_tac "~ (EX act:Acts G. ?P act)" 2);
-by (REPEAT (blast_tac (claset() addSIs [rev_bexI]) 1));
-qed "transient_Join_imp_transient";
-
-Goal "States F = States G  \
-\     ==> (F Join G : transient A) = (F : transient A | G : transient A)";
-by (auto_tac (claset() addSIs [transient_Join_imp_transient], simpset()));
-by (auto_tac (claset(), simpset() addsimps [transient_def, Acts_Join]));
+Goal "F Join G : transient A = \
+\     (F : transient A | G : transient A)";
+by (auto_tac (claset(),
+	      simpset() addsimps [bex_Un, transient_def,
+				  Join_def]));
 qed "transient_Join";
 
-Goal "[| i : I;  eqStates I F |]  \
-\     ==> (JN i:I. F i) : ensures A B = \
-\         ((ALL i:I. F i : constrains (A-B) (A Un B)) & \
-\          (EX i:I. F i : ensures A B))";
+Goal "i : I ==> \
+\     (JN i:I. F i) : ensures A B = \
+\     ((ALL i:I. F i : constrains (A-B) (A Un B)) & \
+\      (EX i:I. F i : ensures A B))";
 by (auto_tac (claset(),
 	      simpset() addsimps [ensures_def, constrains_JN, transient_JN]));
 qed "ensures_JN";
 
 Goalw [ensures_def]
-     "States F = States G  \
-\     ==> F Join G : ensures A B =     \
-\         (F : constrains (A-B) (A Un B) & \
-\          G : constrains (A-B) (A Un B) & \
-\          (F : ensures A B | G : ensures A B))";
+     "F Join G : ensures A B =     \
+\     (F : constrains (A-B) (A Un B) & \
+\      G : constrains (A-B) (A Un B) & \
+\      (F : ensures A B | G : ensures A B))";
 by (auto_tac (claset(),
 	      simpset() addsimps [constrains_Join, transient_Join]));
 qed "ensures_Join";
 
-Goal "[| F : stable A;  G : constrains A A';  \
-\        States F = States G; A <= States G |] \
+Goalw [stable_def, constrains_def, Join_def]
+    "[| F : stable A;  G : constrains A A' |] \
 \    ==> F Join G : constrains A A'";
-by (forward_tac [constrains_imp_subset] 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [stable_def, constrains_def,
-					   ball_Un, Acts_Join]) 1);
+by (asm_full_simp_tac (simpset() addsimps [ball_Un]) 1);
 by (Blast_tac 1);
 qed "stable_constrains_Join";
 
 (*Premise for G cannot use Invariant because  Stable F A  is weaker than
   G : stable A *)
-Goal "[| F : stable A;  G : invariant A;  \
-\        States F = States G |] ==> F Join G : Invariant A";
-by (asm_full_simp_tac (simpset() addsimps [Invariant_def, invariant_def, 
-					   Stable_eq_stable, stable_Join]) 1);
+Goal "[| F : stable A;  G : invariant A |] ==> F Join G : Invariant A";
+by (full_simp_tac (simpset() addsimps [Invariant_def, invariant_def, 
+				       Stable_eq_stable, stable_Join]) 1);
 by (force_tac(claset() addIs [stable_reachable, stable_Int],
 	      simpset() addsimps [Acts_Join]) 1);
 qed "stable_Join_Invariant";
 
-Goal "[| F : stable A;  G : ensures A B;  \
-\        States F = States G |] ==> F Join G : ensures A B";
+Goal "[| F : stable A;  G : ensures A B |] ==> F Join G : ensures A B";
 by (asm_simp_tac (simpset() addsimps [ensures_Join]) 1);
 by (asm_full_simp_tac (simpset() addsimps [stable_def, ensures_def]) 1);
 by (etac constrains_weaken 1);
@@ -361,61 +280,35 @@
 qed "ensures_stable_Join1";
 
 (*As above, but exchanging the roles of F and G*)
-Goal "[| F : ensures A B;  G : stable A;  \
-\        States F = States G |] ==> F Join G : ensures A B";
+Goal "[| F : ensures A B;  G : stable A |] ==> F Join G : ensures A B";
 by (stac Join_commute 1);
-by (dtac sym 1);
 by (blast_tac (claset() addIs [ensures_stable_Join1]) 1);
 qed "ensures_stable_Join2";
 
 
 (** Diff and localTo **)
 
-
-Goal "States (Diff F acts) = States F";
-by (simp_tac (simpset() addsimps [Diff_def]) 1);
-qed "States_Diff";
-
-Goal "Init (Diff F acts) = Init F";
-by (auto_tac (claset() addIs [impOfSubs Init_subset_States],
-	      simpset() addsimps [Diff_def]));
-qed "Init_Diff";
-
-Goal "Acts (Diff F acts) = insert (diag (States F)) (Acts F - acts)";
-by (subgoal_tac "Acts F - acts <= Pow (States F Times States F)" 1);
-by (blast_tac (claset() addEs [equalityE] 
-                        addDs [impOfSubs Acts_subset_Pow_States]) 2);
-by (auto_tac (claset(), simpset() addsimps [Diff_def]));
-qed "Acts_Diff";
-
-Addsimps [States_Diff, Init_Diff, Acts_Diff];
-
-
-Goalw [Join_def] "States F = States G ==> F Join (Diff G (Acts F)) = F Join G";
+Goalw [Join_def, Diff_def] "F Join (Diff G (Acts F)) = F Join G";
 by (rtac program_equalityI 1);
-by (auto_tac (claset(), simpset() addsimps [insert_absorb]));
+by Auto_tac;
 qed "Join_Diff2";
 
-Goalw [Disjoint_def] "States F = States G ==> Disjoint F (Diff G (Acts F))";
+Goalw [Diff_def, Disjoint_def] "Disjoint F (Diff G (Acts F))";
 by Auto_tac;
 qed "Diff_Disjoint";
 
-Goalw [Disjoint_def] "Disjoint F G ==> States F = States G";
-by Auto_tac;
-qed "Disjoint_States_eq";
-
 Goal "[| F Join G : v localTo F;  Disjoint F G |] \
 \     ==> G : (INT z. stable {s. v s = z})";
 by (asm_full_simp_tac 
-    (simpset() addsimps [localTo_def, Disjoint_def,
+    (simpset() addsimps [localTo_def, Diff_def, Disjoint_def,
 			 Acts_Join, stable_def, constrains_def]) 1);
 by (Blast_tac 1);
 qed "localTo_imp_stable";
 
 Goal "[| F Join G : v localTo F;  (s,s') : act;  \
-\        act : Acts G;  Disjoint F G  |] ==> v s' = v s";
+\        act : Acts G;  Disjoint F G |] ==> v s' = v s";
 by (asm_full_simp_tac 
-    (simpset() addsimps [localTo_def, Disjoint_def,
+    (simpset() addsimps [localTo_def, Diff_def, Disjoint_def,
 			 Acts_Join, stable_def, constrains_def]) 1);
 by (Blast_tac 1);
 qed "localTo_imp_equals";
@@ -434,15 +327,9 @@
 \        F Join G : w localTo F;       \
 \        Disjoint F G |]               \
 \     ==> F Join G : constrains {s. P (v s) (w s)} {s. P' (v s) (w s)}";
-by (auto_tac (claset(), 
-	      simpset() addsimps [constrains_def, 
-				  Disjoint_States_eq RS Acts_Join]));
+by (auto_tac (claset(), simpset() addsimps [constrains_def, Acts_Join]));
 by (REPEAT_FIRST (dtac localTo_imp_equals THEN' REPEAT1 o atac));
-by (subgoal_tac "xa : States F" 1);
-by (force_tac
-    (claset() addSDs [Disjoint_States_eq,impOfSubs Acts_subset_Pow_States], 
-     simpset()) 2);
-by (Force_tac 1);
+by Auto_tac;
 qed "constrains_localTo_constrains2";
 
 Goalw [stable_def]
@@ -472,8 +359,7 @@
 by (rtac ballI 1);
 by (subgoal_tac "F Join G : constrains ({s. v s = k} Int {s. v s <= w s}) \
 \                                      ({s. v s = k} Un {s. v s <= w s})" 1);
-by (asm_simp_tac
-    (simpset() addsimps [Disjoint_States_eq RS constrains_Join]) 2);
+by (asm_simp_tac (simpset() addsimps [constrains_Join]) 2);
 by (blast_tac (claset() addIs [constrains_weaken]) 2);
 by (dtac (constrains_imp_Constrains RS Constrains_Int) 1 THEN etac INT_D 1);
 by (etac Constrains_weaken 2);
--- a/src/HOL/UNITY/Union.thy	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/Union.thy	Mon Mar 01 18:38:43 1999 +0100
@@ -11,24 +11,17 @@
 Union = SubstAx + FP +
 
 constdefs
-  eqStates :: ['a set, 'a => 'b program] => bool
-    "eqStates I F == EX St. ALL i:I. States (F i) = St"
-
   JOIN  :: ['a set, 'a => 'b program] => 'b program
-    "JOIN I F == mk_program (INT i:I. States (F i),
-			     INT i:I. Init (F i),
-			     UN i:I. Acts (F i))"
+    "JOIN I F == mk_program (INT i:I. Init (F i), UN i:I. Acts (F i))"
 
   Join :: ['a program, 'a program] => 'a program      (infixl 65)
-    "F Join G == mk_program (States F Int States G,
-			     Init F Int Init G,
-			     Acts F Un Acts G)"
+    "F Join G == mk_program (Init F Int Init G, Acts F Un Acts G)"
 
-  SKIP :: 'a set => 'a program
-    "SKIP states == mk_program (states, states, {})"
+  SKIP :: 'a program
+    "SKIP == mk_program (UNIV, {})"
 
   Diff :: "['a program, ('a * 'a)set set] => 'a program"
-    "Diff F acts == mk_program (States F, Init F, Acts F - acts)"
+    "Diff F acts == mk_program (Init F, Acts F - acts)"
 
   (*The set of systems that regard "v" as local to F*)
   localTo :: ['a => 'b, 'a program] => 'a program set  (infixl 80)
@@ -36,8 +29,7 @@
 
   (*Two programs with disjoint actions, except for identity actions *)
   Disjoint :: ['a program, 'a program] => bool
-    "Disjoint F G == States F = States G &
-                     Acts F Int Acts G <= {diag (States G)}"
+    "Disjoint F G == Acts F Int Acts G <= {Id}"
 
 syntax
   "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)
--- a/src/HOL/UNITY/WFair.ML	Mon Mar 01 18:37:52 1999 +0100
+++ b/src/HOL/UNITY/WFair.ML	Mon Mar 01 18:38:43 1999 +0100
@@ -8,10 +8,6 @@
 From Misra, "A Logic for Concurrent Programming", 1994
 *)
 
-    Goal "[| x:A;  P(x) |] ==> ? x:A. P(x)";
-    by (Blast_tac 1);
-    qed "rev_bexI";
-
 
 overload_1st_set "WFair.transient";
 overload_1st_set "WFair.ensures";
@@ -112,9 +108,16 @@
 by (blast_tac (claset() addIs [leadsto.Union] addDs prems) 1);
 qed "leadsTo_Union";
 
+val prems = Goalw [leadsTo_def]
+    "(!!A. A : S ==> F : leadsTo (A Int C) B) \
+\    ==> F : leadsTo (Union S Int C) B";
+by (simp_tac (HOL_ss addsimps [Int_Union_Union]) 1);
+by (blast_tac (claset() addIs [leadsto.Union] addDs prems) 1);
+qed "leadsTo_Union_Int";
+
 val prems = Goal
     "(!!i. i : I ==> F : leadsTo (A i) B) ==> F : leadsTo (UN i:I. A i) B";
-by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
+by (stac (Union_image_eq RS sym) 1);
 by (blast_tac (claset() addIs leadsTo_Union::prems) 1);
 qed "leadsTo_UN";
 
@@ -152,10 +155,8 @@
 by (blast_tac (claset() addIs [subset_imp_leadsTo, leadsTo_Trans]) 1);
 qed "leadsTo_weaken_R";
 
-
 Goal "[| F : leadsTo A A'; B<=A |] ==> F : leadsTo B A'";
-by (blast_tac (claset() addIs [leadsTo_Basis, leadsTo_Trans, 
-			       subset_imp_leadsTo]) 1);
+by (blast_tac (claset() addIs [leadsTo_Trans, subset_imp_leadsTo]) 1);
 qed_spec_mp "leadsTo_weaken_L";
 
 (*Distributes over binary unions*)
@@ -190,7 +191,7 @@
 val prems = goal thy
    "(!! i. i:I ==> F : leadsTo (A i) (A' i)) \
 \   ==> F : leadsTo (UN i:I. A i) (UN i:I. A' i)";
-by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
+by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1);
 by (blast_tac (claset() addIs [leadsTo_Union, leadsTo_weaken_R] 
                         addIs prems) 1);
 qed "leadsTo_UN_UN";
@@ -269,8 +270,7 @@
    "[| F : leadsTo A A'; F : stable B |] \
 \   ==> F : leadsTo (A Int B) (A' Int B)";
 by (etac leadsTo_induct 1);
-by (simp_tac (simpset() addsimps [Int_Union_Union]) 3);
-by (blast_tac (claset() addIs [leadsTo_Union]) 3);
+by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3);
 by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
 by (rtac leadsTo_Basis 1);
 by (asm_full_simp_tac
@@ -293,8 +293,7 @@
 Goal "[| F : leadsTo A A'; F : constrains B B' |] \
 \     ==> F : leadsTo (A Int B) ((A' Int B) Un (B' - B))";
 by (etac leadsTo_induct 1);
-by (simp_tac (simpset() addsimps [Int_Union_Union]) 3);
-by (blast_tac (claset() addIs [leadsTo_Union]) 3);
+by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3);
 (*Transitivity case has a delicate argument involving "cancellation"*)
 by (rtac leadsTo_Un_duplicate2 2);
 by (etac leadsTo_cancel_Diff1 2);