1(*  Title:      HOL/TLA/Memory/ProcedureInterface.thy
2    Author:     Stephan Merz, University of Munich
3*)
4
5section \<open>Procedure interface for RPC-Memory components\<close>
6
7theory ProcedureInterface
8imports "HOL-TLA.TLA" RPCMemoryParams
9begin
10
11typedecl ('a,'r) chan
12  (* type of channels with argument type 'a and return type 'r.
13     we model a channel as an array of variables (of type chan)
14     rather than a single array-valued variable because the
15     notation gets a little simpler.
16  *)
17type_synonym ('a,'r) channel =" (PrIds \<Rightarrow> ('a,'r) chan) stfun"
18
19
20(* data-level functions *)
21
22consts
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"
27
28
29(* state functions *)
30
31definition 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))"
33
34definition 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
37
38(* slice through array-valued state function *)
39
40definition slice :: "('a \<Rightarrow> 'b) stfun \<Rightarrow> 'a \<Rightarrow> 'b stfun"
41  where "slice x i s \<equiv> x s i"
42
43syntax
44  "_slice" :: "[lift, 'a] \<Rightarrow> lift"  ("(_!_)" [70,70] 70)
45translations
46  "_slice" \<rightleftharpoons> "CONST slice"
47
48
49(* state predicates *)
50
51definition 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 *)
56
57definition 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
63definition 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
69syntax
70  "_Call" :: "['a, 'b, lift] \<Rightarrow> lift"  ("(Call _ _ _)" [90,90,90] 90)
71  "_Return" :: "['a, 'b, lift] \<Rightarrow> lift"  ("(Return _ _ _)" [90,90,90] 90)
72translations
73  "_Call" \<rightleftharpoons> "CONST ACall"
74  "_Return" \<rightleftharpoons> "CONST AReturn"
75
76
77(* temporal formulas *)
78
79definition 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
84definition LegalCaller :: "('a,'r) channel \<Rightarrow> temporal"
85  where "LegalCaller ch == TEMP (\<forall>p. PLegalCaller ch p)"
86
87definition 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
90definition LegalReturner :: "('a,'r) channel \<Rightarrow> temporal"
91  where "LegalReturner ch == TEMP (\<forall>p. PLegalReturner ch p)"
92
93declare slice_def [simp]
94
95lemmas Procedure_defs = caller_def rtrner_def Calling_def ACall_def AReturn_def
96  PLegalCaller_def LegalCaller_def PLegalReturner_def LegalReturner_def
97
98(* Calls and returns change their subchannel *)
99lemma Call_changed: "\<turnstile> Call ch p v \<longrightarrow> <Call ch p v>_((caller ch)!p)"
100  by (auto simp: angle_def ACall_def caller_def Calling_def)
101
102lemma Return_changed: "\<turnstile> Return ch p v \<longrightarrow> <Return ch p v>_((rtrner ch)!p)"
103  by (auto simp: angle_def AReturn_def rtrner_def Calling_def)
104
105end
106