src/HOL/TLA/Memory/ProcedureInterface.thy
changeset 62145 5b946c81dfbf
parent 60592 c9bd1d902f04
child 62146 324bc1ffba12
--- a/src/HOL/TLA/Memory/ProcedureInterface.thy	Mon Jan 11 21:20:44 2016 +0100
+++ b/src/HOL/TLA/Memory/ProcedureInterface.thy	Mon Jan 11 21:21:02 2016 +0100
@@ -16,65 +16,78 @@
   *)
 type_synonym ('a,'r) channel =" (PrIds \<Rightarrow> ('a,'r) chan) stfun"
 
+
+(* data-level functions *)
+
 consts
-  (* data-level functions *)
   cbit          :: "('a,'r) chan \<Rightarrow> bit"
   rbit          :: "('a,'r) chan \<Rightarrow> bit"
   arg           :: "('a,'r) chan \<Rightarrow> 'a"
   res           :: "('a,'r) chan \<Rightarrow> 'r"
 
-  (* state functions *)
-  caller        :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'a)) stfun"
-  rtrner        :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'r)) stfun"
+
+(* state functions *)
 
-  (* state predicates *)
-  Calling   :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> stpred"
+definition caller :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'a)) stfun"
+  where "caller ch == \<lambda>s p. (cbit (ch s p), arg (ch s p))"
 
-  (* actions *)
-  ACall      :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'a stfun \<Rightarrow> action"
-  AReturn    :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'r stfun \<Rightarrow> action"
+definition rtrner :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'r)) stfun"
+  where "rtrner ch == \<lambda>s p. (rbit (ch s p), res (ch s p))"
+
 
-  (* temporal formulas *)
-  PLegalCaller      :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
-  LegalCaller       :: "('a,'r) channel \<Rightarrow> temporal"
-  PLegalReturner    :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
-  LegalReturner     :: "('a,'r) channel \<Rightarrow> temporal"
+(* slice through array-valued state function *)
 
-  (* slice through array-valued state function *)
+consts
   slice        :: "('a \<Rightarrow> 'b) stfun \<Rightarrow> 'a \<Rightarrow> 'b stfun"
-
 syntax
   "_slice"    :: "[lift, 'a] \<Rightarrow> lift"      ("(_!_)" [70,70] 70)
-
-  "_Call"     :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Call _ _ _)" [90,90,90] 90)
-  "_Return"   :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Return _ _ _)" [90,90,90] 90)
-
 translations
   "_slice"  ==  "CONST slice"
-
-  "_Call"   ==  "CONST ACall"
-  "_Return" ==  "CONST AReturn"
-
 defs
   slice_def:     "(PRED (x!i)) s == x s i"
 
-  caller_def:    "caller ch   == \<lambda>s p. (cbit (ch s p), arg (ch s p))"
-  rtrner_def:    "rtrner ch   == \<lambda>s p. (rbit (ch s p), res (ch s p))"
+
+(* state predicates *)
+
+definition Calling :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> stpred"
+  where "Calling ch p == PRED cbit< ch!p > \<noteq> rbit< ch!p >"
+
+
+(* actions *)
 
-  Calling_def:   "Calling ch p  == PRED cbit< ch!p > \<noteq> rbit< ch!p >"
+consts
+  ACall      :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'a stfun \<Rightarrow> action"
+  AReturn    :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'r stfun \<Rightarrow> action"
+syntax
+  "_Call"     :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Call _ _ _)" [90,90,90] 90)
+  "_Return"   :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Return _ _ _)" [90,90,90] 90)
+translations
+  "_Call"   ==  "CONST ACall"
+  "_Return" ==  "CONST AReturn"
+defs
   Call_def:      "(ACT Call ch p v)   == ACT  \<not> $Calling ch p
                                      \<and> (cbit<ch!p>$ \<noteq> $rbit<ch!p>)
                                      \<and> (arg<ch!p>$ = $v)"
   Return_def:    "(ACT Return ch p v) == ACT  $Calling ch p
                                      \<and> (rbit<ch!p>$ = $cbit<ch!p>)
                                      \<and> (res<ch!p>$ = $v)"
-  PLegalCaller_def:      "PLegalCaller ch p == TEMP
-                             Init(\<not> Calling ch p)
-                             \<and> \<box>[\<exists>a. Call ch p a ]_((caller ch)!p)"
-  LegalCaller_def:       "LegalCaller ch == TEMP (\<forall>p. PLegalCaller ch p)"
-  PLegalReturner_def:    "PLegalReturner ch p == TEMP
-                                \<box>[\<exists>v. Return ch p v ]_((rtrner ch)!p)"
-  LegalReturner_def:     "LegalReturner ch == TEMP (\<forall>p. PLegalReturner ch p)"
+
+
+(* temporal formulas *)
+
+definition PLegalCaller :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
+  where "PLegalCaller ch p == TEMP
+     Init(\<not> Calling ch p)
+     \<and> \<box>[\<exists>a. Call ch p a ]_((caller ch)!p)"
+
+definition LegalCaller :: "('a,'r) channel \<Rightarrow> temporal"
+  where "LegalCaller ch == TEMP (\<forall>p. PLegalCaller ch p)"
+
+definition PLegalReturner :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
+  where "PLegalReturner ch p == TEMP \<box>[\<exists>v. Return ch p v ]_((rtrner ch)!p)"
+
+definition LegalReturner :: "('a,'r) channel \<Rightarrow> temporal"
+  where "LegalReturner ch == TEMP (\<forall>p. PLegalReturner ch p)"
 
 declare slice_def [simp]