src/HOL/UNITY/Union.thy
changeset 5584 aad639e56d4e
parent 5313 1861a564d7e2
child 5596 b29d18d8c4d2
--- a/src/HOL/UNITY/Union.thy	Tue Sep 29 15:58:25 1998 +0200
+++ b/src/HOL/UNITY/Union.thy	Tue Sep 29 15:58:47 1998 +0200
@@ -12,15 +12,15 @@
 
 constdefs
    JOIN  :: ['a set, 'a => 'b program] => 'b program
-    "JOIN I prg == (|Init = INT i:I. Init (prg i),
-	           Acts = UN  i:I. Acts (prg i)|)"
+    "JOIN I prg == (|Init  = INT i:I. Init (prg i),
+	             Acts0 = UN  i:I. Acts (prg i)|)"
 
    Join :: ['a program, 'a program] => 'a program
-    "Join prgF prgG == (|Init = Init prgF Int Init prgG,
-			 Acts = Acts prgF Un Acts prgG|)"
+    "Join prgF prgG == (|Init  = Init prgF Int Init prgG,
+			 Acts0 = Acts prgF Un Acts prgG|)"
 
-   Null :: 'a program
-    "Null == (|Init = UNIV, Acts = {id}|)"
+   SKIP :: 'a program
+    "SKIP == (|Init = UNIV, Acts0 = {}|)"
 
 syntax
   "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)