| author | paulson | 
| Tue, 29 Sep 1998 15:58:47 +0200 | |
| changeset 5584 | aad639e56d4e | 
| parent 5313 | 1861a564d7e2 | 
| child 5596 | b29d18d8c4d2 | 
| permissions | -rw-r--r-- | 
| 5252 | 1  | 
(* Title: HOL/UNITY/Union.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1998 University of Cambridge  | 
|
5  | 
||
6  | 
Unions of programs  | 
|
7  | 
||
8  | 
From Misra's Chapter 5: Asynchronous Compositions of Programs  | 
|
9  | 
*)  | 
|
10  | 
||
11  | 
Union = SubstAx + FP +  | 
|
12  | 
||
13  | 
constdefs  | 
|
| 
5313
 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 
paulson 
parents: 
5259 
diff
changeset
 | 
14  | 
JOIN :: ['a set, 'a => 'b program] => 'b program  | 
| 5584 | 15  | 
"JOIN I prg == (|Init = INT i:I. Init (prg i),  | 
16  | 
Acts0 = UN i:I. Acts (prg i)|)"  | 
|
| 5252 | 17  | 
|
| 5259 | 18  | 
Join :: ['a program, 'a program] => 'a program  | 
| 5584 | 19  | 
"Join prgF prgG == (|Init = Init prgF Int Init prgG,  | 
20  | 
Acts0 = Acts prgF Un Acts prgG|)"  | 
|
| 5252 | 21  | 
|
| 5584 | 22  | 
SKIP :: 'a program  | 
23  | 
    "SKIP == (|Init = UNIV, Acts0 = {}|)"
 | 
|
| 5259 | 24  | 
|
| 
5313
 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 
paulson 
parents: 
5259 
diff
changeset
 | 
25  | 
syntax  | 
| 
 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 
paulson 
parents: 
5259 
diff
changeset
 | 
26  | 
  "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)
 | 
| 
 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 
paulson 
parents: 
5259 
diff
changeset
 | 
27  | 
|
| 
 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 
paulson 
parents: 
5259 
diff
changeset
 | 
28  | 
translations  | 
| 
 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 
paulson 
parents: 
5259 
diff
changeset
 | 
29  | 
"JN x:A. B" == "JOIN A (%x. B)"  | 
| 
 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 
paulson 
parents: 
5259 
diff
changeset
 | 
30  | 
|
| 5252 | 31  | 
end  |