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