| author | nipkow | 
| Mon, 21 Aug 2000 17:54:43 +0200 | |
| changeset 9666 | 3572fc1dbe6b | 
| parent 8128 | 3a5864b465e2 | 
| child 11190 | 44e157622cb2 | 
| 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: 
7364 
diff
changeset
 | 
13  | 
instance  | 
| 
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7364 
diff
changeset
 | 
14  | 
program :: (term)ord  | 
| 5597 | 15  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7364 
diff
changeset
 | 
16  | 
defs  | 
| 5597 | 17  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7364 
diff
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: 
7364 
diff
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: 
5597 
diff
changeset
 | 
21  | 
|
| 
8055
 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 
paulson 
parents: 
7399 
diff
changeset
 | 
22  | 
constdefs  | 
| 
 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 
paulson 
parents: 
7399 
diff
changeset
 | 
23  | 
  preserves :: "('a=>'b) => 'a program set"
 | 
| 
 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 
paulson 
parents: 
7399 
diff
changeset
 | 
24  | 
    "preserves v == INT z. stable {s. v s = z}"
 | 
| 
 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 
paulson 
parents: 
7399 
diff
changeset
 | 
25  | 
|
| 
 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 
paulson 
parents: 
7399 
diff
changeset
 | 
26  | 
funPair :: "['a => 'b, 'a => 'c, 'a] => 'b * 'c"  | 
| 
 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 
paulson 
parents: 
7399 
diff
changeset
 | 
27  | 
"funPair f g == %x. (f x, g x)"  | 
| 
 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 
paulson 
parents: 
7399 
diff
changeset
 | 
28  | 
|
| 5597 | 29  | 
end  |