src/HOL/UNITY/Union.thy
author paulson
Thu, 01 Oct 1998 18:28:18 +0200
changeset 5596 b29d18d8c4d2
parent 5584 aad639e56d4e
child 5611 6957f124ca97
permissions -rw-r--r--
abstype of programs

(*  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 == mk_program (INT i:I. Init (prg i),
			       UN  i:I. Acts (prg i))"

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

   SKIP :: 'a program
    "SKIP == mk_program (UNIV, {})"

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

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

end