9 *) |
9 *) |
10 |
10 |
11 Union = SubstAx + FP + |
11 Union = SubstAx + FP + |
12 |
12 |
13 constdefs |
13 constdefs |
|
14 eqStates :: ['a set, 'a => 'b program] => bool |
|
15 "eqStates I F == EX St. ALL i:I. States (F i) = St" |
|
16 |
14 JOIN :: ['a set, 'a => 'b program] => 'b program |
17 JOIN :: ['a set, 'a => 'b program] => 'b program |
15 "JOIN I F == mk_program (INT i:I. Init (F i), UN i:I. Acts (F i))" |
18 "JOIN I F == mk_program (INT i:I. States (F i), |
|
19 INT i:I. Init (F i), |
|
20 UN i:I. Acts (F i))" |
16 |
21 |
17 Join :: ['a program, 'a program] => 'a program (infixl 65) |
22 Join :: ['a program, 'a program] => 'a program (infixl 65) |
18 "F Join G == mk_program (Init F Int Init G, Acts F Un Acts G)" |
23 "F Join G == mk_program (States F Int States G, |
|
24 Init F Int Init G, |
|
25 Acts F Un Acts G)" |
19 |
26 |
20 SKIP :: 'a program |
27 SKIP :: 'a set => 'a program |
21 "SKIP == mk_program (UNIV, {})" |
28 "SKIP states == mk_program (states, states, {})" |
22 |
29 |
23 Diff :: "['a program, ('a * 'a)set set] => 'a program" |
30 Diff :: "['a program, ('a * 'a)set set] => 'a program" |
24 "Diff F acts == mk_program (Init F, Acts F - acts)" |
31 "Diff F acts == mk_program (States F, Init F, Acts F - acts)" |
25 |
32 |
26 (*The set of systems that regard "v" as local to F*) |
33 (*The set of systems that regard "v" as local to F*) |
27 localTo :: ['a => 'b, 'a program] => 'a program set (infixl 80) |
34 localTo :: ['a => 'b, 'a program] => 'a program set (infixl 80) |
28 "v localTo F == {G. ALL z. Diff G (Acts F) : stable {s. v s = z}}" |
35 "v localTo F == {G. ALL z. Diff G (Acts F) : stable {s. v s = z}}" |
29 |
36 |
30 (*Two programs with disjoint actions, except for Id (idling)*) |
37 (*Two programs with disjoint actions, except for identity actions *) |
31 Disjoint :: ['a program, 'a program] => bool |
38 Disjoint :: ['a program, 'a program] => bool |
32 "Disjoint F G == Acts F Int Acts G <= {Id}" |
39 "Disjoint F G == States F = States G & |
|
40 Acts F Int Acts G <= {diag (States G)}" |
33 |
41 |
34 syntax |
42 syntax |
35 "@JOIN" :: [pttrn, 'a set, 'b set] => 'b set ("(3JN _:_./ _)" 10) |
43 "@JOIN" :: [pttrn, 'a set, 'b set] => 'b set ("(3JN _:_./ _)" 10) |
36 |
44 |
37 translations |
45 translations |