| author | wenzelm | 
| Thu, 21 Sep 2000 18:47:18 +0200 | |
| changeset 10054 | 0afe7d951447 | 
| parent 9685 | 6d123a7e30bd | 
| child 10064 | 1a77667b21ef | 
| 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 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 23 | syntax | 
| 7359 | 24 |   "@JOIN1"     :: [pttrns, 'b set] => 'b set         ("(3JN _./ _)" 10)
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 25 |   "@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 | 26 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 27 | translations | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 28 | "JN x:A. B" == "JOIN A (%x. B)" | 
| 7359 | 29 | "JN x y. B" == "JN x. JN y. B" | 
| 30 | "JN x. B" == "JOIN UNIV (%x. B)" | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 31 | |
| 9685 | 32 | syntax (symbols) | 
| 33 |   SKIP      :: 'a program                              ("\\<bottom>")
 | |
| 34 | "op Join" :: ['a program, 'a program] => 'a program (infixl "\\<squnion>" 65) | |
| 35 |   "@JOIN1"  :: [pttrns, 'b set] => 'b set              ("(3\\<Squnion> _./ _)" 10)
 | |
| 36 |   "@JOIN"   :: [pttrn, 'a set, 'b set] => 'b set       ("(3\\<Squnion> _:_./ _)" 10)
 | |
| 37 | ||
| 5252 | 38 | end |