src/HOL/UNITY/Union.thy
author wenzelm
Mon, 10 Aug 1998 17:06:02 +0200
changeset 5290 b755c7240348
parent 5259 86d80749453f
child 5313 1861a564d7e2
permissions -rw-r--r--
suffix, unsuffix moved to Pure/library.ML;

(*  Title:      HOL/UNITY/Union.thy
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1998  University of Cambridge

Unions of programs

From Misra's Chapter 5: Asynchronous Compositions of Programs
*)

Union = SubstAx + FP +

constdefs

   Join :: ['a program, 'a program] => 'a program
    "Join prgF prgG == (|Init = Init prgF Int Init prgG,
			 Acts = Acts prgF Un Acts prgG|)"

   Null :: 'a program
    "Null == (|Init = UNIV, Acts = {id}|)"

end