src/HOL/UNITY/Union.thy
author paulson
Thu, 13 Aug 1998 18:06:40 +0200
changeset 5313 1861a564d7e2
parent 5259 86d80749453f
child 5584 aad639e56d4e
permissions -rw-r--r--
Constrains, Stable, Invariant...more of the substitution axiom, but Union does not work well with them

(*  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 set, 'a => 'b program] => 'b program
    "JOIN I prg == (|Init = INT i:I. Init (prg i),
	           Acts = 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|)"

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

syntax
  "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)

translations
  "JN x:A. B"   == "JOIN A (%x. B)"

end