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