| author | wenzelm | 
| Fri, 22 Oct 1999 20:25:19 +0200 | |
| changeset 7922 | b284079cd902 | 
| parent 7399 | cf780c2bcccf | 
| child 8055 | bb15396278fb | 
| permissions | -rw-r--r-- | 
| 5597 | 1 | (* Title: HOL/UNITY/Comp.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | Composition | |
| 7 | ||
| 8 | From Chandy and Sanders, "Reasoning About Program Composition" | |
| 9 | *) | |
| 10 | ||
| 11 | Comp = Union + | |
| 12 | ||
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7364diff
changeset | 13 | instance | 
| 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7364diff
changeset | 14 | program :: (term)ord | 
| 5597 | 15 | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7364diff
changeset | 16 | defs | 
| 5597 | 17 | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7364diff
changeset | 18 | component_def "F <= H == EX G. F Join G = H" | 
| 5597 | 19 | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7364diff
changeset | 20 | strict_component_def "(F < (H::'a program)) == (F <= H & F ~= H)" | 
| 5612 
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
 paulson parents: 
5597diff
changeset | 21 | |
| 5597 | 22 | end |