| author | oheimb |
| Fri, 11 Dec 1998 18:56:30 +0100 | |
| changeset 6027 | 9dd06eeda95c |
| parent 6012 | 1894bfc4aee9 |
| child 6138 | b7e6e607bb4d |
| 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 |
||
13 |
constdefs |
|
14 |
||
15 |
(*Existential and Universal properties. I formalize the two-program |
|
16 |
case, proving equivalence with Chandy and Sanders's n-ary definitions*) |
|
17 |
||
18 |
ex_prop :: 'a program set => bool |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
19 |
"ex_prop X == |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
20 |
ALL F G. (F:X | G: X) & States F = States G --> (F Join G) : X" |
| 5597 | 21 |
|
22 |
strict_ex_prop :: 'a program set => bool |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
23 |
"strict_ex_prop X == |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
24 |
ALL F G. States F = States G --> (F:X | G: X) = (F Join G : X)" |
| 5597 | 25 |
|
26 |
uv_prop :: 'a program set => bool |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
27 |
"uv_prop X == |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
28 |
SKIP UNIV : X & |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
29 |
(ALL F G. F:X & G: X & States F = States G --> (F Join G) : X)" |
| 5597 | 30 |
|
31 |
strict_uv_prop :: 'a program set => bool |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
32 |
"strict_uv_prop X == |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
33 |
SKIP UNIV : X & |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
34 |
(ALL F G. States F = States G --> (F:X & G: X) = (F Join G : X))" |
| 5597 | 35 |
|
|
5612
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
36 |
(*Ill-defined programs can arise through "Join"*) |
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
37 |
welldef :: 'a program set |
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
38 |
"welldef == {F. Init F ~= {}}"
|
| 5597 | 39 |
|
40 |
component :: ['a program, 'a program] => bool |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
41 |
"component F H == EX G. F Join G = H & States F = States G" |
| 5597 | 42 |
|
43 |
guarantees :: ['a program set, 'a program set] => 'a program set (infixl 65) |
|
44 |
"X guarantees Y == {F. ALL H. component F H --> H:X --> H:Y}"
|
|
45 |
||
|
5612
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
46 |
refines :: ['a program, 'a program, 'a program set] => bool |
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
47 |
("(3_ refines _ wrt _)" [10,10,10] 10)
|
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
48 |
"G refines F wrt X == |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
49 |
States F = States G & |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
50 |
(ALL H. States F = States H & (F Join H) : welldef Int X |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5612
diff
changeset
|
51 |
--> G Join H : welldef Int X)" |
|
5612
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
52 |
|
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
53 |
iso_refines :: ['a program, 'a program, 'a program set] => bool |
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
54 |
("(3_ iso'_refines _ wrt _)" [10,10,10] 10)
|
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
55 |
"G iso_refines F wrt X == |
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
56 |
F : welldef Int X --> G : welldef Int X" |
|
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents:
5597
diff
changeset
|
57 |
|
| 5597 | 58 |
end |