| author | paulson | 
| Thu, 17 Jun 1999 10:35:01 +0200 | |
| changeset 6833 | 15d6c121d75f | 
| parent 6295 | 351b3c2b0d83 | 
| child 7359 | 98a2afab3f86 | 
| 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 | ||
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5648diff
changeset | 8 | Partly from Misra's Chapter 5: Asynchronous Compositions of Programs | 
| 5252 | 9 | *) | 
| 10 | ||
| 11 | Union = SubstAx + FP + | |
| 12 | ||
| 13 | constdefs | |
| 5648 | 14 | JOIN :: ['a set, 'a => 'b program] => 'b program | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 15 | "JOIN I F == mk_program (INT i:I. Init (F i), UN i:I. Acts (F i))" | 
| 5252 | 16 | |
| 5648 | 17 | Join :: ['a program, 'a program] => 'a program (infixl 65) | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 18 | "F Join G == mk_program (Init F Int Init G, Acts F Un Acts G)" | 
| 5252 | 19 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 20 | SKIP :: 'a program | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 21 |     "SKIP == mk_program (UNIV, {})"
 | 
| 5259 | 22 | |
| 5648 | 23 |   Diff :: "['a program, ('a * 'a)set set] => 'a program"
 | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 24 | "Diff F acts == mk_program (Init F, Acts F - acts)" | 
| 5648 | 25 | |
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5648diff
changeset | 26 | (*The set of systems that regard "v" as local to F*) | 
| 5648 | 27 | localTo :: ['a => 'b, 'a program] => 'a program set (infixl 80) | 
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5648diff
changeset | 28 |     "v localTo F == {G. ALL z. Diff G (Acts F) : stable {s. v s = z}}"
 | 
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5648diff
changeset | 29 | |
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5804diff
changeset | 30 | (*Two programs with disjoint actions, except for identity actions *) | 
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5648diff
changeset | 31 | Disjoint :: ['a program, 'a program] => bool | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 32 |     "Disjoint F G == Acts F Int Acts G <= {Id}"
 | 
| 5648 | 33 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 34 | syntax | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 35 |   "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)
 | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 36 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 37 | translations | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 38 | "JN x:A. B" == "JOIN A (%x. B)" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 39 | |
| 5252 | 40 | end |