| 3807 |      1 | (*
 | 
|  |      2 |     File:        RPC.thy
 | 
| 17309 |      3 |     ID:          $Id$
 | 
| 3807 |      4 |     Author:      Stephan Merz
 | 
|  |      5 |     Copyright:   1997 University of Munich
 | 
| 21624 |      6 | *)
 | 
| 3807 |      7 | 
 | 
| 21624 |      8 | header {* RPC-Memory example: RPC specification *}
 | 
| 3807 |      9 | 
 | 
| 17309 |     10 | theory RPC
 | 
|  |     11 | imports RPCParameters ProcedureInterface Memory
 | 
|  |     12 | begin
 | 
| 3807 |     13 | 
 | 
|  |     14 | types
 | 
| 6255 |     15 |   rpcSndChType  = "(rpcOp,Vals) channel"
 | 
|  |     16 |   rpcRcvChType  = "memChType"
 | 
| 3807 |     17 |   rpcStType     = "(PrIds => rpcState) stfun"
 | 
|  |     18 | 
 | 
|  |     19 | consts
 | 
|  |     20 |   (* state predicates *)
 | 
|  |     21 |   RPCInit      :: "rpcRcvChType => rpcStType => PrIds => stpred"
 | 
|  |     22 | 
 | 
|  |     23 |   (* actions *)
 | 
|  |     24 |   RPCFwd     :: "rpcSndChType => rpcRcvChType => rpcStType => PrIds => action"
 | 
|  |     25 |   RPCReject  :: "rpcSndChType => rpcRcvChType => rpcStType => PrIds => action"
 | 
|  |     26 |   RPCFail    :: "rpcSndChType => rpcRcvChType => rpcStType => PrIds => action"
 | 
|  |     27 |   RPCReply   :: "rpcSndChType => rpcRcvChType => rpcStType => PrIds => action"
 | 
|  |     28 |   RPCNext    :: "rpcSndChType => rpcRcvChType => rpcStType => PrIds => action"
 | 
|  |     29 | 
 | 
|  |     30 |   (* temporal *)
 | 
|  |     31 |   RPCIPSpec   :: "rpcSndChType => rpcRcvChType => rpcStType => PrIds => temporal"
 | 
|  |     32 |   RPCISpec   :: "rpcSndChType => rpcRcvChType => rpcStType => temporal"
 | 
|  |     33 | 
 | 
| 17309 |     34 | defs
 | 
|  |     35 |   RPCInit_def:       "RPCInit rcv rst p == PRED ((rst!p = #rpcA) & ~Calling rcv p)"
 | 
| 3807 |     36 | 
 | 
| 17309 |     37 |   RPCFwd_def:        "RPCFwd send rcv rst p == ACT
 | 
| 3807 |     38 |                          $(Calling send p)
 | 
| 6255 |     39 |                          & $(rst!p) = # rpcA
 | 
|  |     40 |                          & IsLegalRcvArg<arg<$(send!p)>>
 | 
|  |     41 |                          & Call rcv p RPCRelayArg<arg<send!p>>
 | 
|  |     42 |                          & (rst!p)$ = # rpcB
 | 
|  |     43 |                          & unchanged (rtrner send!p)"
 | 
| 3807 |     44 | 
 | 
| 17309 |     45 |   RPCReject_def:     "RPCReject send rcv rst p == ACT
 | 
| 6255 |     46 |                            $(rst!p) = # rpcA
 | 
|  |     47 |                          & ~IsLegalRcvArg<arg<$(send!p)>>
 | 
|  |     48 |                          & Return send p #BadCall
 | 
|  |     49 |                          & unchanged ((rst!p), (caller rcv!p))"
 | 
| 3807 |     50 | 
 | 
| 17309 |     51 |   RPCFail_def:       "RPCFail send rcv rst p == ACT
 | 
| 6255 |     52 |                            ~$(Calling rcv p)
 | 
|  |     53 |                          & Return send p #RPCFailure
 | 
|  |     54 |                          & (rst!p)$ = #rpcA
 | 
|  |     55 |                          & unchanged (caller rcv!p)"
 | 
| 3807 |     56 | 
 | 
| 17309 |     57 |   RPCReply_def:      "RPCReply send rcv rst p == ACT
 | 
| 6255 |     58 |                            ~$(Calling rcv p)
 | 
|  |     59 |                          & $(rst!p) = #rpcB
 | 
| 17309 |     60 |                          & Return send p res<rcv!p>
 | 
| 6255 |     61 |                          & (rst!p)$ = #rpcA
 | 
|  |     62 |                          & unchanged (caller rcv!p)"
 | 
| 3807 |     63 | 
 | 
| 17309 |     64 |   RPCNext_def:       "RPCNext send rcv rst p == ACT
 | 
| 6255 |     65 |                         (  RPCFwd send rcv rst p
 | 
|  |     66 |                          | RPCReject send rcv rst p
 | 
|  |     67 |                          | RPCFail send rcv rst p
 | 
|  |     68 |                          | RPCReply send rcv rst p)"
 | 
| 3807 |     69 | 
 | 
| 17309 |     70 |   RPCIPSpec_def:     "RPCIPSpec send rcv rst p == TEMP
 | 
| 6255 |     71 |                            Init RPCInit rcv rst p
 | 
|  |     72 |                          & [][ RPCNext send rcv rst p ]_(rst!p, rtrner send!p, caller rcv!p)
 | 
|  |     73 |                          & WF(RPCNext send rcv rst p)_(rst!p, rtrner send!p, caller rcv!p)"
 | 
| 3807 |     74 | 
 | 
| 17309 |     75 |   RPCISpec_def:      "RPCISpec send rcv rst == TEMP (ALL p. RPCIPSpec send rcv rst p)"
 | 
|  |     76 | 
 | 
| 21624 |     77 | 
 | 
|  |     78 | lemmas RPC_action_defs =
 | 
|  |     79 |   RPCInit_def RPCFwd_def RPCReject_def RPCFail_def RPCReply_def RPCNext_def
 | 
|  |     80 | 
 | 
|  |     81 | lemmas RPC_temp_defs = RPCIPSpec_def RPCISpec_def
 | 
|  |     82 | 
 | 
|  |     83 | 
 | 
|  |     84 | (* The RPC component engages in an action for process p only if there is an outstanding,
 | 
|  |     85 |    unanswered call for that process.
 | 
|  |     86 | *)
 | 
|  |     87 | 
 | 
|  |     88 | lemma RPCidle: "|- ~$(Calling send p) --> ~RPCNext send rcv rst p"
 | 
|  |     89 |   by (auto simp: Return_def RPC_action_defs)
 | 
|  |     90 | 
 | 
|  |     91 | lemma RPCbusy: "|- $(Calling rcv p) & $(rst!p) = #rpcB --> ~RPCNext send rcv rst p"
 | 
|  |     92 |   by (auto simp: RPC_action_defs)
 | 
|  |     93 | 
 | 
|  |     94 | (* RPC failure actions are visible. *)
 | 
|  |     95 | lemma RPCFail_vis: "|- RPCFail send rcv rst p -->  
 | 
|  |     96 |     <RPCNext send rcv rst p>_(rst!p, rtrner send!p, caller rcv!p)"
 | 
|  |     97 |   by (auto dest!: Return_changed [temp_use] simp: angle_def RPCNext_def RPCFail_def)
 | 
|  |     98 | 
 | 
|  |     99 | lemma RPCFail_Next_enabled: "|- Enabled (RPCFail send rcv rst p) -->  
 | 
|  |    100 |     Enabled (<RPCNext send rcv rst p>_(rst!p, rtrner send!p, caller rcv!p))"
 | 
|  |    101 |   by (force elim!: enabled_mono [temp_use] RPCFail_vis [temp_use])
 | 
|  |    102 | 
 | 
|  |    103 | (* Enabledness of some actions *)
 | 
|  |    104 | lemma RPCFail_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==>  
 | 
|  |    105 |     |- ~Calling rcv p & Calling send p --> Enabled (RPCFail send rcv rst p)"
 | 
| 26342 |    106 |   by (tactic {* action_simp_tac (@{simpset} addsimps [thm "RPCFail_def",
 | 
| 21624 |    107 |     thm "Return_def", thm "caller_def", thm "rtrner_def"]) [exI]
 | 
|  |    108 |     [thm "base_enabled", thm "Pair_inject"] 1 *})
 | 
|  |    109 | 
 | 
|  |    110 | lemma RPCReply_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==>  
 | 
|  |    111 |       |- ~Calling rcv p & Calling send p & rst!p = #rpcB  
 | 
|  |    112 |          --> Enabled (RPCReply send rcv rst p)"
 | 
| 26342 |    113 |   by (tactic {* action_simp_tac (@{simpset} addsimps [thm "RPCReply_def",
 | 
| 21624 |    114 |     thm "Return_def", thm "caller_def", thm "rtrner_def"]) [exI]
 | 
|  |    115 |     [thm "base_enabled", thm "Pair_inject"] 1 *})
 | 
| 3807 |    116 | 
 | 
|  |    117 | end
 |