41589
|
1 |
(* Title: HOL/TLA/Memory/MemClerk.thy
|
|
2 |
Author: Stephan Merz, University of Munich
|
21624
|
3 |
*)
|
3807
|
4 |
|
21624
|
5 |
header {* RPC-Memory example: specification of the memory clerk *}
|
3807
|
6 |
|
17309
|
7 |
theory MemClerk
|
|
8 |
imports Memory RPC MemClerkParameters
|
|
9 |
begin
|
3807
|
10 |
|
|
11 |
types
|
|
12 |
(* The clerk takes the same arguments as the memory and sends requests to the RPC *)
|
|
13 |
mClkSndChType = "memChType"
|
|
14 |
mClkRcvChType = "rpcSndChType"
|
|
15 |
mClkStType = "(PrIds => mClkState) stfun"
|
|
16 |
|
36866
|
17 |
definition
|
3807
|
18 |
(* state predicates *)
|
|
19 |
MClkInit :: "mClkRcvChType => mClkStType => PrIds => stpred"
|
36866
|
20 |
where "MClkInit rcv cst p = PRED (cst!p = #clkA & ~Calling rcv p)"
|
3807
|
21 |
|
36866
|
22 |
definition
|
3807
|
23 |
(* actions *)
|
|
24 |
MClkFwd :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action"
|
36866
|
25 |
where "MClkFwd send rcv cst p = ACT
|
6255
|
26 |
$Calling send p
|
|
27 |
& $(cst!p) = #clkA
|
|
28 |
& Call rcv p MClkRelayArg<arg<send!p>>
|
|
29 |
& (cst!p)$ = #clkB
|
|
30 |
& unchanged (rtrner send!p)"
|
|
31 |
|
36866
|
32 |
definition
|
3807
|
33 |
MClkRetry :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action"
|
36866
|
34 |
where "MClkRetry send rcv cst p = ACT
|
6255
|
35 |
$(cst!p) = #clkB
|
|
36 |
& res<$(rcv!p)> = #RPCFailure
|
|
37 |
& Call rcv p MClkRelayArg<arg<send!p>>
|
|
38 |
& unchanged (cst!p, rtrner send!p)"
|
|
39 |
|
36866
|
40 |
definition
|
3807
|
41 |
MClkReply :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action"
|
36866
|
42 |
where "MClkReply send rcv cst p = ACT
|
6255
|
43 |
~$Calling rcv p
|
|
44 |
& $(cst!p) = #clkB
|
|
45 |
& Return send p MClkReplyVal<res<rcv!p>>
|
|
46 |
& (cst!p)$ = #clkA
|
|
47 |
& unchanged (caller rcv!p)"
|
|
48 |
|
36866
|
49 |
definition
|
3807
|
50 |
MClkNext :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action"
|
36866
|
51 |
where "MClkNext send rcv cst p = ACT
|
6255
|
52 |
( MClkFwd send rcv cst p
|
|
53 |
| MClkRetry send rcv cst p
|
|
54 |
| MClkReply send rcv cst p)"
|
|
55 |
|
36866
|
56 |
definition
|
3807
|
57 |
(* temporal *)
|
|
58 |
MClkIPSpec :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => temporal"
|
36866
|
59 |
where "MClkIPSpec send rcv cst p = TEMP
|
6255
|
60 |
Init MClkInit rcv cst p
|
|
61 |
& [][ MClkNext send rcv cst p ]_(cst!p, rtrner send!p, caller rcv!p)
|
|
62 |
& WF(MClkFwd send rcv cst p)_(cst!p, rtrner send!p, caller rcv!p)
|
|
63 |
& SF(MClkReply send rcv cst p)_(cst!p, rtrner send!p, caller rcv!p)"
|
3807
|
64 |
|
36866
|
65 |
definition
|
6255
|
66 |
MClkISpec :: "mClkSndChType => mClkRcvChType => mClkStType => temporal"
|
36866
|
67 |
where "MClkISpec send rcv cst = TEMP (ALL p. MClkIPSpec send rcv cst p)"
|
3807
|
68 |
|
21624
|
69 |
lemmas MC_action_defs =
|
|
70 |
MClkInit_def MClkFwd_def MClkRetry_def MClkReply_def MClkNext_def
|
|
71 |
|
|
72 |
lemmas MC_temp_defs = MClkIPSpec_def MClkISpec_def
|
|
73 |
|
|
74 |
(* The Clerk engages in an action for process p only if there is an outstanding,
|
|
75 |
unanswered call for that process.
|
|
76 |
*)
|
|
77 |
lemma MClkidle: "|- ~$Calling send p & $(cst!p) = #clkA --> ~MClkNext send rcv cst p"
|
|
78 |
by (auto simp: Return_def MC_action_defs)
|
|
79 |
|
|
80 |
lemma MClkbusy: "|- $Calling rcv p --> ~MClkNext send rcv cst p"
|
|
81 |
by (auto simp: Call_def MC_action_defs)
|
|
82 |
|
|
83 |
|
|
84 |
(* Enabledness of actions *)
|
|
85 |
|
|
86 |
lemma MClkFwd_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==>
|
|
87 |
|- Calling send p & ~Calling rcv p & cst!p = #clkA
|
|
88 |
--> Enabled (MClkFwd send rcv cst p)"
|
26359
|
89 |
by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm MClkFwd_def},
|
|
90 |
@{thm Call_def}, @{thm caller_def}, @{thm rtrner_def}]) [exI]
|
|
91 |
[@{thm base_enabled}, @{thm Pair_inject}] 1 *})
|
21624
|
92 |
|
|
93 |
lemma MClkFwd_ch_enabled: "|- Enabled (MClkFwd send rcv cst p) -->
|
|
94 |
Enabled (<MClkFwd send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p))"
|
|
95 |
by (auto elim!: enabled_mono simp: angle_def MClkFwd_def)
|
|
96 |
|
|
97 |
lemma MClkReply_change: "|- MClkReply send rcv cst p -->
|
|
98 |
<MClkReply send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p)"
|
|
99 |
by (auto simp: angle_def MClkReply_def elim: Return_changed [temp_use])
|
|
100 |
|
|
101 |
lemma MClkReply_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==>
|
|
102 |
|- Calling send p & ~Calling rcv p & cst!p = #clkB
|
|
103 |
--> Enabled (<MClkReply send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p))"
|
26342
|
104 |
apply (tactic {* action_simp_tac @{simpset}
|
26359
|
105 |
[@{thm MClkReply_change} RSN (2, @{thm enabled_mono})] [] 1 *})
|
26342
|
106 |
apply (tactic {* action_simp_tac (@{simpset} addsimps
|
26359
|
107 |
[@{thm MClkReply_def}, @{thm Return_def}, @{thm caller_def}, @{thm rtrner_def}])
|
|
108 |
[exI] [@{thm base_enabled}, @{thm Pair_inject}] 1 *})
|
21624
|
109 |
done
|
|
110 |
|
|
111 |
lemma MClkReplyNotRetry: "|- MClkReply send rcv cst p --> ~MClkRetry send rcv cst p"
|
|
112 |
by (auto simp: MClkReply_def MClkRetry_def)
|
17309
|
113 |
|
3807
|
114 |
end
|