author | wenzelm |
Sat, 17 Dec 2005 01:00:40 +0100 | |
changeset 18428 | 4059413acbc1 |
parent 17309 | c43ed29bd197 |
child 21624 | 6f79647cf536 |
permissions | -rw-r--r-- |
3807 | 1 |
(* |
2 |
File: MemClerk.thy |
|
17309 | 3 |
ID: $Id$ |
3807 | 4 |
Author: Stephan Merz |
5 |
Copyright: 1997 University of Munich |
|
6 |
||
7 |
Theory Name: MemClerk |
|
8 |
Logic Image: TLA |
|
9 |
||
10 |
RPC-Memory example: specification of the memory clerk. |
|
11 |
*) |
|
12 |
||
17309 | 13 |
theory MemClerk |
14 |
imports Memory RPC MemClerkParameters |
|
15 |
begin |
|
3807 | 16 |
|
17 |
types |
|
18 |
(* The clerk takes the same arguments as the memory and sends requests to the RPC *) |
|
19 |
mClkSndChType = "memChType" |
|
20 |
mClkRcvChType = "rpcSndChType" |
|
21 |
mClkStType = "(PrIds => mClkState) stfun" |
|
22 |
||
6255 | 23 |
constdefs |
3807 | 24 |
(* state predicates *) |
25 |
MClkInit :: "mClkRcvChType => mClkStType => PrIds => stpred" |
|
6255 | 26 |
"MClkInit rcv cst p == PRED (cst!p = #clkA & ~Calling rcv p)" |
3807 | 27 |
|
28 |
(* actions *) |
|
29 |
MClkFwd :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" |
|
6255 | 30 |
"MClkFwd send rcv cst p == ACT |
31 |
$Calling send p |
|
32 |
& $(cst!p) = #clkA |
|
33 |
& Call rcv p MClkRelayArg<arg<send!p>> |
|
34 |
& (cst!p)$ = #clkB |
|
35 |
& unchanged (rtrner send!p)" |
|
36 |
||
3807 | 37 |
MClkRetry :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" |
6255 | 38 |
"MClkRetry send rcv cst p == ACT |
39 |
$(cst!p) = #clkB |
|
40 |
& res<$(rcv!p)> = #RPCFailure |
|
41 |
& Call rcv p MClkRelayArg<arg<send!p>> |
|
42 |
& unchanged (cst!p, rtrner send!p)" |
|
43 |
||
3807 | 44 |
MClkReply :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" |
6255 | 45 |
"MClkReply send rcv cst p == ACT |
46 |
~$Calling rcv p |
|
47 |
& $(cst!p) = #clkB |
|
48 |
& Return send p MClkReplyVal<res<rcv!p>> |
|
49 |
& (cst!p)$ = #clkA |
|
50 |
& unchanged (caller rcv!p)" |
|
51 |
||
3807 | 52 |
MClkNext :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" |
6255 | 53 |
"MClkNext send rcv cst p == ACT |
54 |
( MClkFwd send rcv cst p |
|
55 |
| MClkRetry send rcv cst p |
|
56 |
| MClkReply send rcv cst p)" |
|
57 |
||
3807 | 58 |
|
59 |
(* temporal *) |
|
60 |
MClkIPSpec :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => temporal" |
|
6255 | 61 |
"MClkIPSpec send rcv cst p == TEMP |
62 |
Init MClkInit rcv cst p |
|
63 |
& [][ MClkNext send rcv cst p ]_(cst!p, rtrner send!p, caller rcv!p) |
|
64 |
& WF(MClkFwd send rcv cst p)_(cst!p, rtrner send!p, caller rcv!p) |
|
65 |
& SF(MClkReply send rcv cst p)_(cst!p, rtrner send!p, caller rcv!p)" |
|
3807 | 66 |
|
6255 | 67 |
MClkISpec :: "mClkSndChType => mClkRcvChType => mClkStType => temporal" |
9517
f58863b1406a
tuned version by Stephan Merz (unbatchified etc.);
wenzelm
parents:
6255
diff
changeset
|
68 |
"MClkISpec send rcv cst == TEMP (ALL p. MClkIPSpec send rcv cst p)" |
3807 | 69 |
|
17309 | 70 |
ML {* use_legacy_bindings (the_context ()) *} |
71 |
||
3807 | 72 |
end |