| author | wenzelm |
| Mon, 10 Dec 2018 20:20:24 +0100 | |
| changeset 69441 | 0bd51c6aaa8b |
| parent 66453 | cc19f7ca2ed6 |
| child 80914 | d97fdabd9e2b |
| permissions | -rw-r--r-- |
| 41589 | 1 |
(* Title: HOL/TLA/Memory/ProcedureInterface.thy |
2 |
Author: Stephan Merz, University of Munich |
|
| 21624 | 3 |
*) |
| 3807 | 4 |
|
| 60592 | 5 |
section \<open>Procedure interface for RPC-Memory components\<close> |
| 3807 | 6 |
|
| 17309 | 7 |
theory ProcedureInterface |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
62146
diff
changeset
|
8 |
imports "HOL-TLA.TLA" RPCMemoryParams |
| 17309 | 9 |
begin |
| 3807 | 10 |
|
| 42018 | 11 |
typedecl ('a,'r) chan
|
| 3807 | 12 |
(* type of channels with argument type 'a and return type 'r. |
| 17309 | 13 |
we model a channel as an array of variables (of type chan) |
14 |
rather than a single array-valued variable because the |
|
| 3807 | 15 |
notation gets a little simpler. |
16 |
*) |
|
| 60588 | 17 |
type_synonym ('a,'r) channel =" (PrIds \<Rightarrow> ('a,'r) chan) stfun"
|
| 3807 | 18 |
|
| 62145 | 19 |
|
20 |
(* data-level functions *) |
|
21 |
||
| 3807 | 22 |
consts |
| 60588 | 23 |
cbit :: "('a,'r) chan \<Rightarrow> bit"
|
24 |
rbit :: "('a,'r) chan \<Rightarrow> bit"
|
|
25 |
arg :: "('a,'r) chan \<Rightarrow> 'a"
|
|
26 |
res :: "('a,'r) chan \<Rightarrow> 'r"
|
|
| 3807 | 27 |
|
| 62145 | 28 |
|
29 |
(* state functions *) |
|
| 3807 | 30 |
|
| 62145 | 31 |
definition caller :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'a)) stfun"
|
32 |
where "caller ch == \<lambda>s p. (cbit (ch s p), arg (ch s p))" |
|
| 3807 | 33 |
|
| 62145 | 34 |
definition rtrner :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'r)) stfun"
|
35 |
where "rtrner ch == \<lambda>s p. (rbit (ch s p), res (ch s p))" |
|
36 |
||
| 3807 | 37 |
|
| 62145 | 38 |
(* slice through array-valued state function *) |
| 3807 | 39 |
|
| 62146 | 40 |
definition slice :: "('a \<Rightarrow> 'b) stfun \<Rightarrow> 'a \<Rightarrow> 'b stfun"
|
41 |
where "slice x i s \<equiv> x s i" |
|
42 |
||
| 6255 | 43 |
syntax |
| 62146 | 44 |
"_slice" :: "[lift, 'a] \<Rightarrow> lift" ("(_!_)" [70,70] 70)
|
| 6255 | 45 |
translations |
| 62146 | 46 |
"_slice" \<rightleftharpoons> "CONST slice" |
| 3807 | 47 |
|
| 62145 | 48 |
|
49 |
(* state predicates *) |
|
50 |
||
51 |
definition Calling :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> stpred"
|
|
52 |
where "Calling ch p == PRED cbit< ch!p > \<noteq> rbit< ch!p >" |
|
53 |
||
54 |
||
55 |
(* actions *) |
|
| 3807 | 56 |
|
| 62146 | 57 |
definition ACall :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'a stfun \<Rightarrow> action"
|
58 |
where "ACall ch p v \<equiv> ACT |
|
59 |
\<not> $Calling ch p |
|
60 |
\<and> (cbit<ch!p>$ \<noteq> $rbit<ch!p>) |
|
61 |
\<and> (arg<ch!p>$ = $v)" |
|
62 |
||
63 |
definition AReturn :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'r stfun \<Rightarrow> action"
|
|
64 |
where "AReturn ch p v == ACT |
|
65 |
$Calling ch p |
|
66 |
\<and> (rbit<ch!p>$ = $cbit<ch!p>) |
|
67 |
\<and> (res<ch!p>$ = $v)" |
|
68 |
||
| 62145 | 69 |
syntax |
| 62146 | 70 |
"_Call" :: "['a, 'b, lift] \<Rightarrow> lift" ("(Call _ _ _)" [90,90,90] 90)
|
71 |
"_Return" :: "['a, 'b, lift] \<Rightarrow> lift" ("(Return _ _ _)" [90,90,90] 90)
|
|
| 62145 | 72 |
translations |
| 62146 | 73 |
"_Call" \<rightleftharpoons> "CONST ACall" |
74 |
"_Return" \<rightleftharpoons> "CONST AReturn" |
|
| 62145 | 75 |
|
76 |
||
77 |
(* temporal formulas *) |
|
78 |
||
79 |
definition PLegalCaller :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
|
|
80 |
where "PLegalCaller ch p == TEMP |
|
81 |
Init(\<not> Calling ch p) |
|
82 |
\<and> \<box>[\<exists>a. Call ch p a ]_((caller ch)!p)" |
|
83 |
||
84 |
definition LegalCaller :: "('a,'r) channel \<Rightarrow> temporal"
|
|
85 |
where "LegalCaller ch == TEMP (\<forall>p. PLegalCaller ch p)" |
|
86 |
||
87 |
definition PLegalReturner :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
|
|
88 |
where "PLegalReturner ch p == TEMP \<box>[\<exists>v. Return ch p v ]_((rtrner ch)!p)" |
|
89 |
||
90 |
definition LegalReturner :: "('a,'r) channel \<Rightarrow> temporal"
|
|
91 |
where "LegalReturner ch == TEMP (\<forall>p. PLegalReturner ch p)" |
|
| 17309 | 92 |
|
| 21624 | 93 |
declare slice_def [simp] |
94 |
||
| 62146 | 95 |
lemmas Procedure_defs = caller_def rtrner_def Calling_def ACall_def AReturn_def |
| 21624 | 96 |
PLegalCaller_def LegalCaller_def PLegalReturner_def LegalReturner_def |
97 |
||
98 |
(* Calls and returns change their subchannel *) |
|
| 60588 | 99 |
lemma Call_changed: "\<turnstile> Call ch p v \<longrightarrow> <Call ch p v>_((caller ch)!p)" |
| 62146 | 100 |
by (auto simp: angle_def ACall_def caller_def Calling_def) |
| 21624 | 101 |
|
| 60588 | 102 |
lemma Return_changed: "\<turnstile> Return ch p v \<longrightarrow> <Return ch p v>_((rtrner ch)!p)" |
| 62146 | 103 |
by (auto simp: angle_def AReturn_def rtrner_def Calling_def) |
| 3807 | 104 |
|
105 |
end |