| 3807 |      1 | (*
 | 
|  |      2 |     File:        MemClerkParameters.thy
 | 
|  |      3 |     Author:      Stephan Merz
 | 
|  |      4 |     Copyright:   1997 University of Munich
 | 
|  |      5 | 
 | 
|  |      6 |     Theory Name: MemClerkParameters
 | 
|  |      7 |     Logic Image: TLA
 | 
|  |      8 | 
 | 
|  |      9 |     RPC-Memory example: Parameters of the memory clerk.
 | 
|  |     10 | *)
 | 
|  |     11 | 
 | 
|  |     12 | MemClerkParameters = RPCParameters + 
 | 
|  |     13 | 
 | 
|  |     14 | datatype  mClkState  =  clkA | clkB
 | 
|  |     15 | 
 | 
|  |     16 | types
 | 
|  |     17 |   (* types sent on the clerk's send and receive channels are argument types
 | 
|  |     18 |      of the memory and the RPC, respectively *)
 | 
| 6255 |     19 |   mClkSndArgType   = "memOp"
 | 
|  |     20 |   mClkRcvArgType   = "rpcOp"
 | 
| 3807 |     21 | 
 | 
| 6255 |     22 | constdefs
 | 
| 3807 |     23 |   (* translate a memory call to an RPC call *)
 | 
| 6255 |     24 |   MClkRelayArg     :: "memOp => rpcOp"
 | 
|  |     25 |     "MClkRelayArg marg == memcall marg"
 | 
| 3807 |     26 |   (* translate RPC failures to memory failures *)
 | 
|  |     27 |   MClkReplyVal     :: "Vals => Vals"
 | 
| 6255 |     28 |     "MClkReplyVal v == if v = RPCFailure then MemFailure else v"
 | 
| 3807 |     29 | 
 | 
|  |     30 | end
 | 
|  |     31 | 
 |