src/HOL/UNITY/Union.thy
changeset 6012 1894bfc4aee9
parent 5804 8e0a4c4fd67b
child 6295 351b3c2b0d83
--- a/src/HOL/UNITY/Union.thy	Wed Dec 02 16:14:09 1998 +0100
+++ b/src/HOL/UNITY/Union.thy	Thu Dec 03 10:45:06 1998 +0100
@@ -11,25 +11,33 @@
 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. Init (F i), UN i:I. Acts (F i))"
+    "JOIN I F == mk_program (INT i:I. States (F i),
+			     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 (Init F Int Init G, Acts F Un Acts G)"
+    "F Join G == mk_program (States F Int States G,
+			     Init F Int Init G,
+			     Acts F Un Acts G)"
 
-  SKIP :: 'a program
-    "SKIP == mk_program (UNIV, {})"
+  SKIP :: 'a set => 'a program
+    "SKIP states == mk_program (states, states, {})"
 
   Diff :: "['a program, ('a * 'a)set set] => 'a program"
-    "Diff F acts == mk_program (Init F, Acts F - acts)"
+    "Diff F acts == mk_program (States F, 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)
     "v localTo F == {G. ALL z. Diff G (Acts F) : stable {s. v s = z}}"
 
-  (*Two programs with disjoint actions, except for Id (idling)*)
+  (*Two programs with disjoint actions, except for identity actions *)
   Disjoint :: ['a program, 'a program] => bool
-    "Disjoint F G == Acts F Int Acts G <= {Id}"
+    "Disjoint F G == States F = States G &
+                     Acts F Int Acts G <= {diag (States G)}"
 
 syntax
   "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)