src/HOL/UNITY/Union.thy
author paulson
Thu, 03 Dec 1998 10:45:06 +0100
changeset 6012 1894bfc4aee9
parent 5804 8e0a4c4fd67b
child 6295 351b3c2b0d83
permissions -rw-r--r--
Addition of the States component; parts of Comp not working

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

Unions of programs

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

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. 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 (States F Int States G,
			     Init F Int Init G,
			     Acts F Un Acts G)"

  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 (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 identity actions *)
  Disjoint :: ['a program, 'a program] => bool
    "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)

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

end