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