src/HOL/TLA/Memory/ProcedureInterface.thy
changeset 62145 5b946c81dfbf
parent 60592 c9bd1d902f04
child 62146 324bc1ffba12
equal deleted inserted replaced
62144:bdab93ccfaf8 62145:5b946c81dfbf
    14      rather than a single array-valued variable because the
    14      rather than a single array-valued variable because the
    15      notation gets a little simpler.
    15      notation gets a little simpler.
    16   *)
    16   *)
    17 type_synonym ('a,'r) channel =" (PrIds \<Rightarrow> ('a,'r) chan) stfun"
    17 type_synonym ('a,'r) channel =" (PrIds \<Rightarrow> ('a,'r) chan) stfun"
    18 
    18 
       
    19 
       
    20 (* data-level functions *)
       
    21 
    19 consts
    22 consts
    20   (* data-level functions *)
       
    21   cbit          :: "('a,'r) chan \<Rightarrow> bit"
    23   cbit          :: "('a,'r) chan \<Rightarrow> bit"
    22   rbit          :: "('a,'r) chan \<Rightarrow> bit"
    24   rbit          :: "('a,'r) chan \<Rightarrow> bit"
    23   arg           :: "('a,'r) chan \<Rightarrow> 'a"
    25   arg           :: "('a,'r) chan \<Rightarrow> 'a"
    24   res           :: "('a,'r) chan \<Rightarrow> 'r"
    26   res           :: "('a,'r) chan \<Rightarrow> 'r"
    25 
    27 
    26   (* state functions *)
       
    27   caller        :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'a)) stfun"
       
    28   rtrner        :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'r)) stfun"
       
    29 
    28 
    30   (* state predicates *)
    29 (* state functions *)
    31   Calling   :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> stpred"
       
    32 
    30 
    33   (* actions *)
    31 definition caller :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'a)) stfun"
    34   ACall      :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'a stfun \<Rightarrow> action"
    32   where "caller ch == \<lambda>s p. (cbit (ch s p), arg (ch s p))"
    35   AReturn    :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'r stfun \<Rightarrow> action"
       
    36 
    33 
    37   (* temporal formulas *)
    34 definition rtrner :: "('a,'r) channel \<Rightarrow> (PrIds \<Rightarrow> (bit * 'r)) stfun"
    38   PLegalCaller      :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
    35   where "rtrner ch == \<lambda>s p. (rbit (ch s p), res (ch s p))"
    39   LegalCaller       :: "('a,'r) channel \<Rightarrow> temporal"
       
    40   PLegalReturner    :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
       
    41   LegalReturner     :: "('a,'r) channel \<Rightarrow> temporal"
       
    42 
    36 
    43   (* slice through array-valued state function *)
    37 
       
    38 (* slice through array-valued state function *)
       
    39 
       
    40 consts
    44   slice        :: "('a \<Rightarrow> 'b) stfun \<Rightarrow> 'a \<Rightarrow> 'b stfun"
    41   slice        :: "('a \<Rightarrow> 'b) stfun \<Rightarrow> 'a \<Rightarrow> 'b stfun"
    45 
       
    46 syntax
    42 syntax
    47   "_slice"    :: "[lift, 'a] \<Rightarrow> lift"      ("(_!_)" [70,70] 70)
    43   "_slice"    :: "[lift, 'a] \<Rightarrow> lift"      ("(_!_)" [70,70] 70)
    48 
       
    49   "_Call"     :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Call _ _ _)" [90,90,90] 90)
       
    50   "_Return"   :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Return _ _ _)" [90,90,90] 90)
       
    51 
       
    52 translations
    44 translations
    53   "_slice"  ==  "CONST slice"
    45   "_slice"  ==  "CONST slice"
    54 
       
    55   "_Call"   ==  "CONST ACall"
       
    56   "_Return" ==  "CONST AReturn"
       
    57 
       
    58 defs
    46 defs
    59   slice_def:     "(PRED (x!i)) s == x s i"
    47   slice_def:     "(PRED (x!i)) s == x s i"
    60 
    48 
    61   caller_def:    "caller ch   == \<lambda>s p. (cbit (ch s p), arg (ch s p))"
       
    62   rtrner_def:    "rtrner ch   == \<lambda>s p. (rbit (ch s p), res (ch s p))"
       
    63 
    49 
    64   Calling_def:   "Calling ch p  == PRED cbit< ch!p > \<noteq> rbit< ch!p >"
    50 (* state predicates *)
       
    51 
       
    52 definition Calling :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> stpred"
       
    53   where "Calling ch p == PRED cbit< ch!p > \<noteq> rbit< ch!p >"
       
    54 
       
    55 
       
    56 (* actions *)
       
    57 
       
    58 consts
       
    59   ACall      :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'a stfun \<Rightarrow> action"
       
    60   AReturn    :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> 'r stfun \<Rightarrow> action"
       
    61 syntax
       
    62   "_Call"     :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Call _ _ _)" [90,90,90] 90)
       
    63   "_Return"   :: "['a, 'b, lift] \<Rightarrow> lift"    ("(Return _ _ _)" [90,90,90] 90)
       
    64 translations
       
    65   "_Call"   ==  "CONST ACall"
       
    66   "_Return" ==  "CONST AReturn"
       
    67 defs
    65   Call_def:      "(ACT Call ch p v)   == ACT  \<not> $Calling ch p
    68   Call_def:      "(ACT Call ch p v)   == ACT  \<not> $Calling ch p
    66                                      \<and> (cbit<ch!p>$ \<noteq> $rbit<ch!p>)
    69                                      \<and> (cbit<ch!p>$ \<noteq> $rbit<ch!p>)
    67                                      \<and> (arg<ch!p>$ = $v)"
    70                                      \<and> (arg<ch!p>$ = $v)"
    68   Return_def:    "(ACT Return ch p v) == ACT  $Calling ch p
    71   Return_def:    "(ACT Return ch p v) == ACT  $Calling ch p
    69                                      \<and> (rbit<ch!p>$ = $cbit<ch!p>)
    72                                      \<and> (rbit<ch!p>$ = $cbit<ch!p>)
    70                                      \<and> (res<ch!p>$ = $v)"
    73                                      \<and> (res<ch!p>$ = $v)"
    71   PLegalCaller_def:      "PLegalCaller ch p == TEMP
    74 
    72                              Init(\<not> Calling ch p)
    75 
    73                              \<and> \<box>[\<exists>a. Call ch p a ]_((caller ch)!p)"
    76 (* temporal formulas *)
    74   LegalCaller_def:       "LegalCaller ch == TEMP (\<forall>p. PLegalCaller ch p)"
    77 
    75   PLegalReturner_def:    "PLegalReturner ch p == TEMP
    78 definition PLegalCaller :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
    76                                 \<box>[\<exists>v. Return ch p v ]_((rtrner ch)!p)"
    79   where "PLegalCaller ch p == TEMP
    77   LegalReturner_def:     "LegalReturner ch == TEMP (\<forall>p. PLegalReturner ch p)"
    80      Init(\<not> Calling ch p)
       
    81      \<and> \<box>[\<exists>a. Call ch p a ]_((caller ch)!p)"
       
    82 
       
    83 definition LegalCaller :: "('a,'r) channel \<Rightarrow> temporal"
       
    84   where "LegalCaller ch == TEMP (\<forall>p. PLegalCaller ch p)"
       
    85 
       
    86 definition PLegalReturner :: "('a,'r) channel \<Rightarrow> PrIds \<Rightarrow> temporal"
       
    87   where "PLegalReturner ch p == TEMP \<box>[\<exists>v. Return ch p v ]_((rtrner ch)!p)"
       
    88 
       
    89 definition LegalReturner :: "('a,'r) channel \<Rightarrow> temporal"
       
    90   where "LegalReturner ch == TEMP (\<forall>p. PLegalReturner ch p)"
    78 
    91 
    79 declare slice_def [simp]
    92 declare slice_def [simp]
    80 
    93 
    81 lemmas Procedure_defs = caller_def rtrner_def Calling_def Call_def Return_def
    94 lemmas Procedure_defs = caller_def rtrner_def Calling_def Call_def Return_def
    82   PLegalCaller_def LegalCaller_def PLegalReturner_def LegalReturner_def
    95   PLegalCaller_def LegalCaller_def PLegalReturner_def LegalReturner_def