src/HOL/TLA/Memory/MemoryImplementation.thy
changeset 42786 06a38b936342
parent 42772 2acb503fd857
child 42814 5af15f1e2ef6
equal deleted inserted replaced
42785:15ec9b3c14cc 42786:06a38b936342
   789 
   789 
   790 (* Frequently needed abbreviation: distinguish between idling and non-idling
   790 (* Frequently needed abbreviation: distinguish between idling and non-idling
   791    steps of the implementation, and try to solve the idling case by simplification
   791    steps of the implementation, and try to solve the idling case by simplification
   792 *)
   792 *)
   793 ML {*
   793 ML {*
   794 fun split_idle_tac ctxt simps i =
   794 fun split_idle_tac ctxt =
   795   let val ss = simpset_of ctxt in
   795   SELECT_GOAL
   796     TRY (rtac @{thm actionI} i) THEN
   796    (TRY (rtac @{thm actionI} 1) THEN
   797     InductTacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" i THEN
   797     InductTacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" 1 THEN
   798     rewrite_goals_tac @{thms action_rews} THEN
   798     rewrite_goals_tac @{thms action_rews} THEN
   799     forward_tac [temp_use @{thm Step1_4_7}] i THEN
   799     forward_tac [temp_use @{thm Step1_4_7}] 1 THEN
   800     asm_full_simp_tac (ss addsimps simps) i
   800     asm_full_simp_tac (simpset_of ctxt) 1);
   801   end
       
   802 *}
   801 *}
       
   802 
       
   803 method_setup split_idle = {*
       
   804   Method.sections (Simplifier.simp_modifiers @ Splitter.split_modifiers)
       
   805     >> (K (SIMPLE_METHOD' o split_idle_tac))
       
   806 *} ""
       
   807 
   803 (* ----------------------------------------------------------------------
   808 (* ----------------------------------------------------------------------
   804    Combine steps 1.2 and 1.4 to prove that the implementation satisfies
   809    Combine steps 1.2 and 1.4 to prove that the implementation satisfies
   805    the specification's next-state relation.
   810    the specification's next-state relation.
   806 *)
   811 *)
   807 
   812 
   808 (* Steps that leave all variables unchanged are safe, so I may assume
   813 (* Steps that leave all variables unchanged are safe, so I may assume
   809    that some variable changes in the proof that a step is safe. *)
   814    that some variable changes in the proof that a step is safe. *)
   810 lemma unchanged_safe: "|- (~unchanged (e p, c p, r p, m p, rmhist!p)
   815 lemma unchanged_safe: "|- (~unchanged (e p, c p, r p, m p, rmhist!p)
   811              --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p))
   816              --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p))
   812          --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
   817          --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
   813   apply (tactic {* split_idle_tac @{context} [@{thm square_def}] 1 *})
   818   apply (split_idle simp: square_def)
   814   apply force
   819   apply force
   815   done
   820   done
   816 (* turn into (unsafe, looping!) introduction rule *)
   821 (* turn into (unsafe, looping!) introduction rule *)
   817 lemmas unchanged_safeI = impI [THEN unchanged_safe [action_use], standard]
   822 lemmas unchanged_safeI = impI [THEN unchanged_safe [action_use], standard]
   818 
   823 
   880 
   885 
   881 (* ------------------------------ State S1 ------------------------------ *)
   886 (* ------------------------------ State S1 ------------------------------ *)
   882 
   887 
   883 lemma S1_successors: "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   888 lemma S1_successors: "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   884          --> (S1 rmhist p)$ | (S2 rmhist p)$"
   889          --> (S1 rmhist p)$ | (S2 rmhist p)$"
   885   apply (tactic "split_idle_tac @{context} [] 1")
   890   apply split_idle
   886   apply (auto dest!: Step1_2_1 [temp_use])
   891   apply (auto dest!: Step1_2_1 [temp_use])
   887   done
   892   done
   888 
   893 
   889 (* Show that the implementation can satisfy the high-level fairness requirements
   894 (* Show that the implementation can satisfy the high-level fairness requirements
   890    by entering the state S1 infinitely often.
   895    by entering the state S1 infinitely often.
   914 
   919 
   915 (* ------------------------------ State S2 ------------------------------ *)
   920 (* ------------------------------ State S2 ------------------------------ *)
   916 
   921 
   917 lemma S2_successors: "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   922 lemma S2_successors: "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   918          --> (S2 rmhist p)$ | (S3 rmhist p)$"
   923          --> (S2 rmhist p)$ | (S3 rmhist p)$"
   919   apply (tactic "split_idle_tac @{context} [] 1")
   924   apply split_idle
   920   apply (auto dest!: Step1_2_2 [temp_use])
   925   apply (auto dest!: Step1_2_2 [temp_use])
   921   done
   926   done
   922 
   927 
   923 lemma S2MClkFwd_successors: "|- ($S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
   928 lemma S2MClkFwd_successors: "|- ($S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
   924          & <MClkFwd memCh crCh cst p>_(c p)
   929          & <MClkFwd memCh crCh cst p>_(c p)
   940 
   945 
   941 (* ------------------------------ State S3 ------------------------------ *)
   946 (* ------------------------------ State S3 ------------------------------ *)
   942 
   947 
   943 lemma S3_successors: "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   948 lemma S3_successors: "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   944          --> (S3 rmhist p)$ | (S4 rmhist p | S6 rmhist p)$"
   949          --> (S3 rmhist p)$ | (S4 rmhist p | S6 rmhist p)$"
   945   apply (tactic "split_idle_tac @{context} [] 1")
   950   apply split_idle
   946   apply (auto dest!: Step1_2_3 [temp_use])
   951   apply (auto dest!: Step1_2_3 [temp_use])
   947   done
   952   done
   948 
   953 
   949 lemma S3RPC_successors: "|- ($S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
   954 lemma S3RPC_successors: "|- ($S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
   950          & <RPCNext crCh rmCh rst p>_(r p)
   955          & <RPCNext crCh rmCh rst p>_(r p)
   968 (* ------------- State S4 -------------------------------------------------- *)
   973 (* ------------- State S4 -------------------------------------------------- *)
   969 
   974 
   970 lemma S4_successors: "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   975 lemma S4_successors: "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
   971         & (ALL l. $MemInv mm l)
   976         & (ALL l. $MemInv mm l)
   972         --> (S4 rmhist p)$ | (S5 rmhist p)$"
   977         --> (S4 rmhist p)$ | (S5 rmhist p)$"
   973   apply (tactic "split_idle_tac @{context} [] 1")
   978   apply split_idle
   974   apply (auto dest!: Step1_2_4 [temp_use])
   979   apply (auto dest!: Step1_2_4 [temp_use])
   975   done
   980   done
   976 
   981 
   977 (* --------- State S4a: S4 /\ (ires p = NotAResult) ------------------------ *)
   982 (* --------- State S4a: S4 /\ (ires p = NotAResult) ------------------------ *)
   978 
   983 
   979 lemma S4a_successors: "|- $(S4 rmhist p & ires!p = #NotAResult)
   984 lemma S4a_successors: "|- $(S4 rmhist p & ires!p = #NotAResult)
   980          & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l)
   985          & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l)
   981          --> (S4 rmhist p & ires!p = #NotAResult)$
   986          --> (S4 rmhist p & ires!p = #NotAResult)$
   982              | ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$"
   987              | ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$"
   983   apply (tactic {* split_idle_tac @{context} [@{thm m_def}] 1 *})
   988   apply split_idle
   984   apply (auto dest!: Step1_2_4 [temp_use])
   989   apply (auto dest!: Step1_2_4 [temp_use])
   985   done
   990   done
   986 
   991 
   987 lemma S4aRNext_successors: "|- ($(S4 rmhist p & ires!p = #NotAResult)
   992 lemma S4aRNext_successors: "|- ($(S4 rmhist p & ires!p = #NotAResult)
   988          & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l))
   993          & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l))
  1009 (* ---------- State S4b: S4 /\ (ires p # NotAResult) --------------------------- *)
  1014 (* ---------- State S4b: S4 /\ (ires p # NotAResult) --------------------------- *)
  1010 
  1015 
  1011 lemma S4b_successors: "|- $(S4 rmhist p & ires!p ~= #NotAResult)
  1016 lemma S4b_successors: "|- $(S4 rmhist p & ires!p ~= #NotAResult)
  1012          & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l)
  1017          & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l)
  1013          --> (S4 rmhist p & ires!p ~= #NotAResult)$ | (S5 rmhist p)$"
  1018          --> (S4 rmhist p & ires!p ~= #NotAResult)$ | (S5 rmhist p)$"
  1014   apply (tactic {* split_idle_tac @{context} [@{thm m_def}] 1 *})
  1019   apply (split_idle simp: m_def)
  1015   apply (auto dest!: WriteResult [temp_use] Step1_2_4 [temp_use] ReadResult [temp_use])
  1020   apply (auto dest!: WriteResult [temp_use] Step1_2_4 [temp_use] ReadResult [temp_use])
  1016   done
  1021   done
  1017 
  1022 
  1018 lemma S4bReturn_successors: "|- ($(S4 rmhist p & ires!p ~= #NotAResult)
  1023 lemma S4bReturn_successors: "|- ($(S4 rmhist p & ires!p ~= #NotAResult)
  1019          & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
  1024          & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
  1038 
  1043 
  1039 (* ------------------------------ State S5 ------------------------------ *)
  1044 (* ------------------------------ State S5 ------------------------------ *)
  1040 
  1045 
  1041 lemma S5_successors: "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
  1046 lemma S5_successors: "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
  1042          --> (S5 rmhist p)$ | (S6 rmhist p)$"
  1047          --> (S5 rmhist p)$ | (S6 rmhist p)$"
  1043   apply (tactic "split_idle_tac @{context} [] 1")
  1048   apply split_idle
  1044   apply (auto dest!: Step1_2_5 [temp_use])
  1049   apply (auto dest!: Step1_2_5 [temp_use])
  1045   done
  1050   done
  1046 
  1051 
  1047 lemma S5RPC_successors: "|- ($S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
  1052 lemma S5RPC_successors: "|- ($S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
  1048          & <RPCNext crCh rmCh rst p>_(r p)
  1053          & <RPCNext crCh rmCh rst p>_(r p)
  1064 
  1069 
  1065 (* ------------------------------ State S6 ------------------------------ *)
  1070 (* ------------------------------ State S6 ------------------------------ *)
  1066 
  1071 
  1067 lemma S6_successors: "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
  1072 lemma S6_successors: "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)
  1068          --> (S1 rmhist p)$ | (S3 rmhist p)$ | (S6 rmhist p)$"
  1073          --> (S1 rmhist p)$ | (S3 rmhist p)$ | (S6 rmhist p)$"
  1069   apply (tactic "split_idle_tac @{context} [] 1")
  1074   apply split_idle
  1070   apply (auto dest!: Step1_2_6 [temp_use])
  1075   apply (auto dest!: Step1_2_6 [temp_use])
  1071   done
  1076   done
  1072 
  1077 
  1073 lemma S6MClkReply_successors:
  1078 lemma S6MClkReply_successors:
  1074   "|- ($S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
  1079   "|- ($S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))
  1279          & $ImpInv rmhist p
  1284          & $ImpInv rmhist p
  1280          --> (S4 rmhist p)$ & unchanged (e p, c p, r p, rmhist!p)"
  1285          --> (S4 rmhist p)$ & unchanged (e p, c p, r p, rmhist!p)"
  1281   apply clarsimp
  1286   apply clarsimp
  1282   apply (drule WriteS4 [action_use])
  1287   apply (drule WriteS4 [action_use])
  1283    apply assumption
  1288    apply assumption
  1284   apply (tactic "split_idle_tac @{context} [] 1")
  1289   apply split_idle
  1285   apply (auto simp: ImpNext_def dest!: S4EnvUnch [temp_use] S4ClerkUnch [temp_use]
  1290   apply (auto simp: ImpNext_def dest!: S4EnvUnch [temp_use] S4ClerkUnch [temp_use]
  1286     S4RPCUnch [temp_use])
  1291     S4RPCUnch [temp_use])
  1287      apply (auto simp: square_def dest: S4Write [temp_use])
  1292      apply (auto simp: square_def dest: S4Write [temp_use])
  1288   done
  1293   done
  1289 
  1294