diff --git a/thys/Universal_Turing_Machine/BlanksDoNotMatter.thy b/thys/Universal_Turing_Machine/BlanksDoNotMatter.thy --- a/thys/Universal_Turing_Machine/BlanksDoNotMatter.thy +++ b/thys/Universal_Turing_Machine/BlanksDoNotMatter.thy @@ -1,1260 +1,1260 @@ (* Title: thys/BlanksDoNotMatter.thy Author: Franz Regensburger (FABR) 04/2022 *) section \Trailing Blanks on the input tape do not matter\ theory BlanksDoNotMatter imports Turing begin (* ----- Configure sledgehammer ------ *) sledgehammer_params[minimize=false,preplay_timeout=10,timeout=30,strict=true, provers="e z3 cvc4 vampire "] (* sledgehammer_params[minimize=false,preplay_timeout=10,timeout=30,verbose=true,strict=true, provers="spass e z3 cvc4 vampire "] *) (* ----------------------------------- *) subsection \Replication of symbols\ abbreviation exponent :: "'a \ nat \ 'a list" ("_ \ _" [100, 99] 100) where "x \ n == replicate n x" lemma hd_repeat_cases: "P (hd (a \ m @ r)) \ (m = 0 \ P (hd r)) \ (\nat. m = Suc nat \ P a)" by (cases m,auto) lemma hd_repeat_cases': "P (hd (a \ m @ r)) = (if m = 0 then P (hd r) else P a)" by auto (* hd_repeat_cases \ hd_repeat_cases'. * However, hd_repeat_cases' is better for rewriting *) lemma "(if m = 0 then P (hd r) else P a) = ((m = 0 \ P (hd r)) \ (\nat. m = Suc nat \ P a))" proof - have "(if m = 0 then P (hd r) else P a) = P (hd (a \ m @ r))" by auto also have "... = ((m = 0 \ P (hd r)) \ (\nat. m = Suc nat \ P a))" by (simp add: iffI hd_repeat_cases) finally show ?thesis . qed lemma split_head_repeat[simp]: "Oc # list1 = Bk \ j @ list2 \ j = 0 \ Oc # list1 = list2" "Bk # list1 = Oc \ j @ list2 \ j = 0 \ Bk # list1 = list2" "Bk \ j @ list2 = Oc # list1 \ j = 0 \ Oc # list1 = list2" "Oc \ j @ list2 = Bk # list1 \ j = 0 \ Bk # list1 = list2" by(cases j;force)+ lemma Bk_no_Oc_repeatE[elim]: "Bk # list = Oc \ t \ RR" by (cases t, auto) lemma replicate_Suc_1: "a \ (z1 + Suc z2) = (a \ z1) @ (a \ Suc z2)" by (meson replicate_add) lemma replicate_Suc_2: "a \ (z1 + Suc z2) = (a \ Suc z1) @ (a \ z2)" by (simp add: replicate_add) subsection \Trailing blanks on the left tape do not matter\ text\In this section we will show that we may add or remove trailing blanks on the initial left and right portions of the tape at will. However, we may not add or remove trailing blanks on the tape resulting from the computation. The resulting tape is completely determined by the contents of the initial tape.\ (* -------------------- Remove trailing blanks from the left tape -------------------------- *) lemma step_left_tape_ShrinkBkCtx_right_Nil: assumes "step0 (s,CL@Bk\ z1 , []) tm = (s',l',r')" and "za < z1" shows "\CL' zb. l' = CL'@Bk\za@Bk\zb \ (step0 (s,CL@Bk\za, []) tm = (s',CL'@Bk\za,r') \ step0 (s,CL@Bk\za, []) tm = (s',CL'@Bk\(za-1),r'))" proof (cases "fetch tm (s - 0) (read [])") case (Pair a s2) then have A1: "fetch tm (s - 0) (read []) = (a, s2)" . show ?thesis proof (cases a) assume "a = WB" from \a = WB\ and assms A1 have "step0 (s, CL@Bk \ z1, []) tm = (s2, CL@Bk \ z1, [Bk])" by auto moreover from \a = WB\ and assms A1 have "step0 (s, CL@Bk\za, []) tm = (s2,CL@Bk\za , [Bk])" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, []) tm = (s2,CL@Bk\za , [Bk])" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, []) tm = (s2, CL @ Bk \ z1, [Bk])\ assms(1) by auto next assume "a = WO" from \a = WO\ and assms A1 have "step0 (s, CL@Bk \ z1, []) tm = (s2, CL@Bk \ z1, [Oc])" by auto moreover from \a = WO\ and assms A1 have "step0 (s, CL@Bk\za, []) tm = (s2,CL@Bk\za , [Oc])" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, []) tm = (s2,CL@Bk\za , [Oc])" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, []) tm = (s2, CL @ Bk \ z1, [Oc])\ assms(1) by auto next assume "a = Nop" from \a = Nop\ and assms A1 have "step0 (s, CL@Bk \ z1, []) tm = (s2, CL@Bk \ z1, [])" by auto moreover from \a = Nop\ and assms A1 have "step0 (s, CL@Bk\za, []) tm = (s2,CL@Bk\za , [])" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, []) tm = (s2,CL@Bk\za , [])" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, []) tm = (s2, CL @ Bk \ z1, [])\ assms(1) by auto next assume "a = R" from \a = R\ and assms A1 have "step0 (s, CL@Bk \ z1, []) tm = (s2, [Bk]@CL@Bk\z1, [])" by auto moreover from \a = R\ and assms A1 have "step0 (s, CL@Bk\za, []) tm = (s2,[Bk]@CL@Bk\za, [])" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, []) tm = (s2,[Bk]@CL@Bk\za, [])" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, []) tm = (s2, [Bk] @ CL @ Bk \ z1, [])\ by (metis append_Cons append_Nil assms(1) fst_conv snd_conv) next assume "a = L" show ?thesis proof (cases CL) case Nil then have "CL = []" . then show ?thesis proof (cases z1) case 0 then have "z1 = 0" . with assms and \CL = []\ show ?thesis by auto next case (Suc nat) then have "z1 = Suc nat" . from \a = L\ and \CL = []\ and \z1 = Suc nat\ and assms and A1 have "step0 (s, CL@Bk \ z1, []) tm = (s2, []@Bk \(z1-1), [Bk])" by auto moreover from \a = L\ and \CL = []\ and A1 have "step0 (s, CL@Bk\za, []) tm = (s2, []@Bk\(za-1) , [Bk])" by auto ultimately have "Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, []) tm = (s2, []@Bk\(za-1) , [Bk])" using assms using \z1 = Suc nat\ by (metis diff_Suc_1 le_eq_less_or_eq less_Suc_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) with assms and \CL = []\ and \z1 = Suc nat\ and \step0 (s, CL@Bk\z1, []) tm = (s2, [] @ Bk \ (z1 - 1), [Bk])\ show ?thesis by auto qed next case (Cons c cs) then have "CL = c # cs" . from \a = L\ and \CL = c # cs\ and assms and A1 have "step0 (s, CL@Bk \ z1, []) tm = (s2, cs@Bk \z1, [c])" by auto moreover from \a = L\ and \CL = c # cs\ and A1 have "step0 (s, CL@Bk\za, []) tm = (s2, cs@Bk \za, [c])" by auto ultimately have "Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, []) tm = (s2, cs@Bk \za, [c])" using assms by (metis One_nat_def Suc_pred add_diff_inverse_nat neq0_conv not_less_eq not_less_zero replicate_add) with assms and \CL = c # cs\ and \Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, []) tm = (s2, cs@Bk \za, [c])\ show ?thesis using \step0 (s, CL @ Bk \ z1, []) tm = (s2, cs @ Bk \ z1, [c])\ by (metis fst_conv nat_le_linear not_less ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add snd_conv) qed qed qed lemma step_left_tape_ShrinkBkCtx_right_Bk: assumes "step0 (s,CL@Bk\ z1 , Bk#rs) tm = (s',l',r')" and "za < z1" shows "\CL' zb. l' = CL'@Bk\za@Bk\zb \ (step0 (s,CL@Bk\za, Bk#rs) tm = (s',CL'@Bk\za,r') \ step0 (s,CL@Bk\za, Bk#rs) tm = (s',CL'@Bk\(za-1),r'))" proof (cases "fetch tm (s - 0) (read (Bk#rs))") case (Pair a s2) then have A1: "fetch tm (s - 0) (read (Bk#rs)) = (a, s2)" . show ?thesis proof (cases a) assume "a = WB" from \a = WB\ and assms A1 have "step0 (s, CL@Bk \ z1, Bk#rs) tm = (s2, CL@Bk \ z1, Bk#rs)" by (auto simp add: split: if_splits) moreover from \a = WB\ and assms A1 have "step0 (s, CL@Bk\za, Bk#rs) tm = (s2,CL@Bk\za , Bk#rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Bk#rs) tm = (s2,CL@Bk\za , Bk#rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Bk#rs) tm = (s2, CL @ Bk \ z1, Bk#rs)\ assms(1) by auto next assume "a = WO" from \a = WO\ and assms A1 have "step0 (s, CL@Bk \ z1, Bk#rs) tm = (s2, CL@Bk \ z1, Oc#rs)" by auto moreover from \a = WO\ and assms A1 have "step0 (s, CL@Bk\za, Bk#rs) tm = (s2,CL@Bk\za , Oc#rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Bk#rs) tm = (s2,CL@Bk\za , Oc#rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Bk#rs) tm = (s2, CL @ Bk \ z1, Oc#rs)\ assms(1) by auto next assume "a = Nop" from \a = Nop\ and assms A1 have "step0 (s, CL@Bk \ z1, Bk#rs) tm = (s2, CL@Bk \ z1, Bk#rs)" by auto moreover from \a = Nop\ and assms A1 have "step0 (s, CL@Bk\za, Bk#rs) tm = (s2,CL@Bk\za , Bk#rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Bk#rs) tm = (s2,CL@Bk\za , Bk#rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Bk#rs) tm = (s2, CL @ Bk \ z1, Bk#rs)\ assms(1) by auto next assume "a = R" from \a = R\ and assms A1 have "step0 (s, CL@Bk \ z1, Bk#rs) tm = (s2, [Bk]@CL@Bk\z1, rs)" by auto moreover from \a = R\ and assms A1 have "step0 (s, CL@Bk\za, Bk#rs) tm = (s2,[Bk]@CL@Bk\za, rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Bk#rs) tm = (s2,[Bk]@CL@Bk\za, rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Bk#rs) tm = (s2, [Bk] @ CL @ Bk \ z1, rs)\ by (metis append_Cons append_Nil assms(1) fst_conv snd_conv) next assume "a = L" show ?thesis proof (cases CL) case Nil then have "CL = []" . then show ?thesis proof (cases z1) case 0 then have "z1 = 0" . with assms and \CL = []\ show ?thesis by auto next case (Suc nat) then have "z1 = Suc nat" . from \a = L\ and \CL = []\ and \z1 = Suc nat\ and assms and A1 have "step0 (s, CL@Bk \ z1, Bk#rs) tm = (s2, []@Bk \(z1-1), Bk#Bk#rs)" by auto moreover from \a = L\ and \CL = []\ and A1 have "step0 (s, CL@Bk\za, Bk#rs) tm = (s2, []@Bk\(za-1) , Bk#Bk#rs)" by auto ultimately have "Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, Bk#rs) tm = (s2, []@Bk\(za-1) , Bk#Bk#rs)" using assms using \z1 = Suc nat\ by (metis diff_Suc_1 le_eq_less_or_eq less_Suc_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) with assms and \CL = []\ and \z1 = Suc nat\ and \step0 (s, CL@Bk\z1, Bk#rs) tm = (s2, [] @ Bk \ (z1 - 1), Bk#Bk#rs)\ show ?thesis by auto qed next case (Cons c cs) then have "CL = c # cs" . from \a = L\ and \CL = c # cs\ and assms and A1 have "step0 (s, CL@Bk \ z1, Bk#rs) tm = (s2, cs@Bk \z1, c#Bk#rs)" by auto moreover from \a = L\ and \CL = c # cs\ and A1 have "step0 (s, CL@Bk\za, Bk#rs) tm = (s2, cs@Bk \za, c#Bk#rs)" by auto ultimately have "Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, Bk#rs) tm = (s2, cs@Bk \za, c#Bk#rs)" using assms by (metis One_nat_def Suc_pred add_diff_inverse_nat neq0_conv not_less_eq not_less_zero replicate_add) with assms and \CL = c # cs\ and \Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, Bk#rs) tm = (s2, cs@Bk \za, c#Bk#rs)\ show ?thesis using \step0 (s, CL @ Bk \ z1, Bk#rs) tm = (s2, cs @ Bk \ z1, c#Bk#rs)\ by (metis fst_conv nat_le_linear not_less ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add snd_conv) qed qed qed lemma step_left_tape_ShrinkBkCtx_right_Oc: assumes "step0 (s,CL@Bk\ z1 , Oc#rs) tm = (s',l',r')" and "za < z1" shows "\CL' zb. l' = CL'@Bk\za@Bk\zb \ (step0 (s,CL@Bk\za, Oc#rs) tm = (s',CL'@Bk\za,r') \ step0 (s,CL@Bk\za, Oc#rs) tm = (s',CL'@Bk\(za-1),r'))" proof (cases "fetch tm (s - 0) (read (Oc#rs))") case (Pair a s2) then have A1: "fetch tm (s - 0) (read (Oc#rs)) = (a, s2)" . show ?thesis proof (cases a) assume "a = WB" from \a = WB\ and assms A1 have "step0 (s, CL@Bk \ z1, Oc#rs) tm = (s2, CL@Bk \ z1, Bk#rs)" by auto moreover from \a = WB\ and assms A1 have "step0 (s, CL@Bk\za, Oc#rs) tm = (s2,CL@Bk\za , Bk#rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Oc#rs) tm = (s2,CL@Bk\za ,Bk#rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Oc#rs) tm = (s2, CL @ Bk \ z1, Bk#rs)\ assms(1) by auto next assume "a = WO" from \a = WO\ and assms A1 have "step0 (s, CL@Bk \ z1, Oc#rs) tm = (s2, CL@Bk \ z1, Oc#rs)" by auto moreover from \a = WO\ and assms A1 have "step0 (s, CL@Bk\za, Oc#rs) tm = (s2,CL@Bk\za , Oc#rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Oc#rs) tm = (s2,CL@Bk\za , Oc#rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Oc#rs) tm = (s2, CL @ Bk \ z1, Oc#rs)\ assms(1) by auto next assume "a = Nop" from \a = Nop\ and assms A1 have "step0 (s, CL@Bk \ z1, Oc#rs) tm = (s2, CL@Bk \ z1, Oc#rs)" by auto moreover from \a = Nop\ and assms A1 have "step0 (s, CL@Bk\za, Oc#rs) tm = (s2,CL@Bk\za , Oc#rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Oc#rs) tm = (s2,CL@Bk\za , Oc#rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Oc#rs) tm = (s2, CL @ Bk \ z1, Oc#rs)\ assms(1) by auto next assume "a = R" from \a = R\ and assms A1 have "step0 (s, CL@Bk \ z1, Oc#rs) tm = (s2, [Oc]@CL@Bk\z1, rs)" by auto moreover from \a = R\ and assms A1 have "step0 (s, CL@Bk\za, Oc#rs) tm = (s2,[Oc]@CL@Bk\za, rs)" by auto ultimately have "Bk\z1 = Bk\za@Bk\(z1-za) \ step0 (s, CL@Bk\za, Oc#rs) tm = (s2,[Oc]@CL@Bk\za, rs)" using assms by (metis le_eq_less_or_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) then show ?thesis using \step0 (s, CL @ Bk \ z1, Oc#rs) tm = (s2, [Oc] @ CL @ Bk \ z1, rs)\ by (metis append_Cons append_Nil assms(1) fst_conv snd_conv) next assume "a = L" show ?thesis proof (cases CL) case Nil then have "CL = []" . then show ?thesis proof (cases z1) case 0 then have "z1 = 0" . with assms and \CL = []\ show ?thesis by auto next case (Suc nat) then have "z1 = Suc nat" . from \a = L\ and \CL = []\ and \z1 = Suc nat\ and assms and A1 have "step0 (s, CL@Bk \ z1, Oc#rs) tm = (s2, []@Bk \(z1-1), Bk#Oc#rs)" by auto moreover from \a = L\ and \CL = []\ and A1 have "step0 (s, CL@Bk\za, Oc#rs) tm = (s2, []@Bk\(za-1) , Bk#Oc#rs)" by auto ultimately have "Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, Oc#rs) tm = (s2, []@Bk\(za-1) , Bk#Oc#rs)" using assms using \z1 = Suc nat\ by (metis diff_Suc_1 le_eq_less_or_eq less_Suc_eq ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add) with assms and \CL = []\ and \z1 = Suc nat\ and \step0 (s, CL@Bk\z1, Oc#rs) tm = (s2, [] @ Bk \ (z1 - 1), Bk#Oc#rs)\ show ?thesis by auto qed next case (Cons c cs) then have "CL = c # cs" . from \a = L\ and \CL = c # cs\ and assms and A1 have "step0 (s, CL@Bk \ z1, Oc#rs) tm = (s2, cs@Bk \z1, c#Oc#rs)" by auto moreover from \a = L\ and \CL = c # cs\ and A1 have "step0 (s, CL@Bk\za, Oc#rs) tm = (s2, cs@Bk \za, c#Oc#rs)" by auto ultimately have "Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, Oc#rs) tm = (s2, cs@Bk \za, c#Oc#rs)" using assms by (metis One_nat_def Suc_pred add_diff_inverse_nat neq0_conv not_less_eq not_less_zero replicate_add) with assms and \CL = c # cs\ and \Bk\(z1-1) = Bk\za@Bk\(z1-1-za) \ step0 (s, CL@Bk\za, Oc#rs) tm = (s2, cs@Bk \za, c#Oc#rs)\ show ?thesis using \step0 (s, CL @ Bk \ z1, Oc#rs) tm = (s2, cs @ Bk \ z1, c#Oc#rs)\ by (metis fst_conv nat_le_linear not_less ordered_cancel_comm_monoid_diff_class.add_diff_inverse replicate_add snd_conv) qed qed qed corollary step_left_tape_ShrinkBkCtx: assumes "step0 (s,CL@Bk\ z1 , r) tm = (s',l',r')" and "za < z1" shows "\zb CL'. l' = CL'@Bk\za@Bk\zb \ (step0 (s,CL@Bk\za, r) tm = (s',CL'@Bk\za,r') \ step0 (s,CL@Bk\za, r) tm = (s',CL'@Bk\(za-1),r'))" proof (cases r) case Nil then show ?thesis using step_left_tape_ShrinkBkCtx_right_Nil using assms by blast next case (Cons rx rs) then have "r = rx # rs" . show ?thesis proof (cases rx) case Bk with assms and \r = rx # rs\ show ?thesis using step_left_tape_ShrinkBkCtx_right_Bk by blast next case Oc with assms and \r = rx # rs\ show ?thesis using step_left_tape_ShrinkBkCtx_right_Oc by blast qed qed (* propagate the step lemmas step_left_tape_ShrinkBkCtx* to steps *) lemma steps_left_tape_ShrinkBkCtx_arbitrary_CL: "\ steps0 (s, CL@Bk\z1 , r) tm stp = (s',l',r'); 0 < z1 \ \ \zb CL'. l' = CL'@Bk\zb \ steps0 (s,CL, r) tm stp = (s',CL',r')" proof (induct stp arbitrary: s CL z1 r s' l' r' z1) case 0 assume "steps0 (s, CL @ Bk \ z1, r) tm 0 = (s', l', r')" and "0 < z1" then show ?case using less_imp_add_positive replicate_add by fastforce next fix stp s CL z1 r s' l' r' assume IV: "\s2 CL2 z12 r2 s2' l2' r2'. \steps0 (s2, CL2 @ Bk \ z12, r2) tm stp = (s2', l2', r2'); 0 < z12\ \ \zb2' CL2'. l2' = CL2' @ Bk \ zb2' \ steps0 (s2, CL2, r2) tm stp = (s2', CL2', r2')" and major: "steps0 (s, CL @ Bk \ z1, r) tm (Suc stp) = (s', l', r')" and minor: "0< z1" show "\zb CL'. l' = CL' @ Bk \ zb \ steps0 (s, CL, r) tm (Suc stp) = (s', CL', r')" proof - have F1: "steps0 (s, CL, r) tm (Suc stp) = step0 (steps0 (s, CL, r) tm stp) tm" by (rule step_red) have "steps0 (s, CL @ Bk \ z1, r) tm (Suc stp) = step0 (steps0 (s, CL @ Bk \ z1, r) tm stp) tm" by (rule step_red) with major have F3: "step0 (steps0 (s, CL @ Bk \ z1, r) tm stp) tm = (s', l', r')" by auto show ?thesis proof (cases z1) case 0 then have "z1 = 0" . with minor show ?thesis by auto next case (Suc z1') then have "z1 = Suc z1'" . show ?thesis proof (cases "steps0 (s, CL @ Bk \ z1, r) tm stp") case (fields sx lx rx) then have C: "steps0 (s, CL @ Bk \ z1, r) tm stp = (sx, lx, rx)" . with minor and IV have F0: "\zb2' CL2'. lx = CL2' @ Bk \ zb2' \ steps0 (s, CL, r) tm stp = (sx, CL2', rx)" by auto then obtain zb2' CL2' where w_zb2'_CL2'_zc2': "lx = CL2' @ Bk \ zb2' \ steps0 (s, CL, r) tm stp = (sx, CL2', rx)" by blast from F3 and C have "step0 (sx,lx,rx) tm = (s',l',r')" by auto with w_zb2'_CL2'_zc2' have F4: "step0 (sx,CL2' @ Bk \ zb2',rx) tm = (s',l',r')" by auto then have "step0 (sx,CL2' @ Bk \ (zb2'),rx) tm = (s',l',r')" by (simp add: replicate_add) show ?thesis proof (cases zb2') case 0 then show ?thesis using F1 \step0 (sx, lx, rx) tm = (s', l', r')\ append_Nil2 replicate_0 w_zb2'_CL2'_zc2' by auto next case (Suc zb3') then have "zb2' = Suc zb3'" . then show ?thesis by (metis F1 F4 append_Nil2 diff_is_0_eq' replicate_0 self_append_conv2 step_left_tape_ShrinkBkCtx w_zb2'_CL2'_zc2' zero_le_one zero_less_Suc) qed qed qed qed qed (* -------------------- Add trailing blanks on the left tape ------------------------------- *) lemma step_left_tape_EnlargeBkCtx_eq_Bks: assumes "step0 (s,Bk\ z1, r) tm = (s',l',r')" shows "step0 (s,Bk\(z1+Suc z2), r) tm = (s',l'@Bk\Suc z2,r') \ step0 (s,Bk\(z1+Suc z2), r) tm = (s',l'@Bk\z2,r')" proof (cases s) assume "s = 0" with assms have "step0 (s, Bk\(z1+Suc z2), r) tm = (s',l'@Bk\Suc z2,r')" using replicate_Suc_1 by fastforce then show ?thesis by auto next fix s2 assume "s = Suc s2" then show ?thesis proof (cases r) assume "r = []" then show "step0 (s, Bk \ (z1 + Suc z2), r) tm = (s', l' @ Bk \ Suc z2, r') \ step0 (s, Bk \ (z1 + Suc z2), r) tm = (s', l' @ Bk \ z2, r')" proof (cases "fetch tm (s - 0) (read r)") case (Pair a s3) then have "fetch tm (s - 0) (read r) = (a, s3)" . then show ?thesis proof (cases a) case WB from \a = WB\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk\(z1+Suc z2), [Bk])" by auto moreover from \a = WB\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ z1, [Bk])" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce next case WO from \a = WO\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s,Bk\(z1+Suc z2), r) tm = (s3, Bk\(z1+Suc z2), [Oc])" by auto moreover from \a = WO\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ z1, [Oc])" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce next case L from \a = L\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk \ (z1 + z2), [Bk])" by auto moreover from \a = L\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ (z1-1), [Bk])" by auto ultimately show ?thesis by (metis (no_types, lifting) Pair_inject add_Suc_right add_eq_if assms diff_is_0_eq' replicate_add zero_le_one) next case R from \a = R\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk# Bk\(z1+Suc z2), [])" by auto moreover from \a = R\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk# Bk \ z1, [])" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce next case Nop from \a = Nop\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk\(z1+Suc z2), [])" by auto moreover from \a = Nop\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ z1, [])" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce qed qed next fix ra rrs assume "r = ra # rrs" then show "step0 (s, Bk \ (z1 + Suc z2), r) tm = (s', l' @ Bk \ Suc z2, r') \ step0 (s, Bk \ (z1 + Suc z2), r) tm = (s', l' @ Bk \ z2, r')" proof (cases "fetch tm (s - 0) (read r)") case (Pair a s3) then have "fetch tm (s - 0) (read r) = (a, s3)" . then show ?thesis proof (cases a) case WB from \a = WB\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk\(z1+Suc z2), Bk# rrs)" by auto moreover from \a = WB\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ z1, Bk# rrs)" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce next case WO from \a = WO\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk\(z1+Suc z2), Oc# rrs)" by auto moreover from \a = WO\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ z1, Oc# rrs)" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce next case L from \a = L\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk \ (z1 + z2), Bk#ra#rrs)" by auto moreover from \a = L\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ (z1-1), Bk#ra#rrs)" by auto ultimately show ?thesis by (metis (no_types, lifting) Pair_inject add_Suc_right add_eq_if assms diff_is_0_eq' replicate_add zero_le_one) next case R from \a = R\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, ra# Bk\(z1+Suc z2), rrs)" by auto moreover from \a = R\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, ra# Bk \ z1,rrs)" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce next case Nop from \a = Nop\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk\(z1+Suc z2), r) tm = (s3, Bk\(z1+Suc z2), ra # rrs)" by auto moreover from \a = Nop\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, Bk \ z1, r) tm = (s3, Bk \ z1, ra # rrs)" by auto ultimately show ?thesis using assms replicate_Suc_1 by fastforce qed qed qed qed lemma step_left_tape_EnlargeBkCtx_eq_Bk_C_Bks: assumes "step0 (s,(Bk#C)@Bk\ z1, r) tm = (s',l',r')" shows " step0 (s,(Bk#C)@Bk\(z1+z2), r) tm = (s',l'@Bk\z2,r')" proof (cases s) assume "s = 0" with assms show "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s', l' @ Bk \ z2, r')" by (auto simp add: replicate_add) next fix s2 assume "s = Suc s2" then show "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s', l' @ Bk \ z2, r')" proof (cases r) assume "r = []" then show ?thesis proof (cases "fetch tm (s - 0) (read r)") case (Pair a s3) then have "fetch tm (s - 0) (read r) = (a, s3)" . then show ?thesis proof (cases a) case WB from \a = WB\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk # C) @ Bk \ (z1 + z2), [Bk])" by auto moreover from \a = WB\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (Bk # C) @ Bk \ z1, [Bk])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case WO from \a = WO\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk # C) @ Bk \ (z1 + z2), [Oc])" by auto moreover from \a = WO\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (Bk # C) @ Bk \ z1, [Oc])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case L from \a = L\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (C) @ Bk \ (z1 + z2), [Bk])" by auto moreover from \a = L\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (C) @ Bk \ z1, [Bk])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case R from \a = R\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk# Bk # C) @ Bk \ (z1 + z2), [])" by auto moreover from \a = R\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (Bk#Bk # C) @ Bk \ z1, [])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case Nop from \a = Nop\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk # C) @ Bk \ (z1 + z2), [])" by auto moreover from \a = Nop\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (Bk # C) @ Bk \ z1, [])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) qed qed next fix ra rrs assume "r = ra # rrs" then show "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s', l' @ Bk \ z2, r')" proof (cases "fetch tm (s - 0) (read r)") case (Pair a s3) then have "fetch tm (s - 0) (read r) = (a, s3)" . then show ?thesis proof (cases a) case WB from \a = WB\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk # C) @ Bk \ (z1 + z2), Bk# rrs)" by auto moreover from \a = WB\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (Bk # C) @ Bk \ z1, Bk# rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case WO from \a = WO\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk # C) @ Bk \ (z1 + z2), Oc# rrs)" by auto moreover from \a = WO\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (Bk # C) @ Bk \ z1, Oc# rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case L from \a = L\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (C) @ Bk \ (z1 + z2), Bk#ra#rrs)" by auto moreover from \a = L\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (C) @ Bk \ z1, Bk#ra#rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case R from \a = R\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (ra# Bk # C) @ Bk \ (z1 + z2), rrs)" by auto moreover from \a = R\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (ra#Bk # C) @ Bk \ z1,rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case Nop from \a = Nop\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk # C) @ Bk \ (z1 + z2), ra # rrs)" by auto moreover from \a = Nop\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Bk # C) @ Bk \ z1, r) tm = (s3, (Bk # C) @ Bk \ z1, ra # rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) qed qed qed qed lemma step_left_tape_EnlargeBkCtx_eq_Oc_C_Bks: assumes "step0 (s,(Oc#C)@Bk\ z1, r) tm = (s',l',r')" shows " step0 (s,(Oc#C)@Bk\(z1+z2), r) tm = (s',l'@Bk\z2,r')" proof (cases s) assume "s = 0" with assms show "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s', l' @ Bk \ z2, r')" by (auto simp add: replicate_add) next fix s2 assume "s = Suc s2" then show "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s', l' @ Bk \ z2, r')" proof (cases r) assume "r = []" then show ?thesis proof (cases "fetch tm (s - 0) (read r)") case (Pair a s3) then have "fetch tm (s - 0) (read r) = (a, s3)" . then show ?thesis proof (cases a) case WB from \a = WB\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (Oc # C) @ Bk \ (z1 + z2), [Bk])" by auto moreover from \a = WB\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (Oc # C) @ Bk \ z1, [Bk])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case WO from \a = WO\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (Oc # C) @ Bk \ (z1 + z2), [Oc])" by auto moreover from \a = WO\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (Oc # C) @ Bk \ z1, [Oc])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case L from \a = L\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (C) @ Bk \ (z1 + z2), [Oc])" by auto moreover from \a = L\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (C) @ Bk \ z1, [Oc])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case R from \a = R\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (Bk# Oc # C) @ Bk \ (z1 + z2), [])" by auto moreover from \a = R\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (Bk#Oc # C) @ Bk \ z1, [])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case Nop from \a = Nop\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (Oc # C) @ Bk \ (z1 + z2), [])" by auto moreover from \a = Nop\ \r = []\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (Oc # C) @ Bk \ z1, [])" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) qed qed next fix ra rrs assume "r = ra # rrs" then show "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s', l' @ Bk \ z2, r')" proof (cases "fetch tm (s - 0) (read r)") case (Pair a s3) then have "fetch tm (s - 0) (read r) = (a, s3)" . then show ?thesis proof (cases a) case WB from \a = WB\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (Oc # C) @ Bk \ (z1 + z2), Bk# rrs)" by auto moreover from \a = WB\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (Oc # C) @ Bk \ z1, Bk# rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case WO from \a = WO\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (Oc # C) @ Bk \ (z1 + z2), Oc# rrs)" by auto moreover from \a = WO\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (Oc # C) @ Bk \ z1, Oc# rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case L from \a = L\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (C) @ Bk \ (z1 + z2), Oc#ra#rrs)" by auto moreover from \a = L\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (C) @ Bk \ z1, Oc#ra#rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case R from \a = R\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (ra# Oc # C) @ Bk \ (z1 + z2), rrs)" by auto moreover from \a = R\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (ra#Oc # C) @ Bk \ z1,rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) next case Nop from \a = Nop\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ (z1 + z2), r) tm = (s3, (Oc # C) @ Bk \ (z1 + z2), ra # rrs)" by auto moreover from \a = Nop\ \r = ra # rrs\ \fetch tm (s - 0) (read r) = (a, s3)\ have "step0 (s, (Oc # C) @ Bk \ z1, r) tm = (s3, (Oc # C) @ Bk \ z1, ra # rrs)" by auto ultimately show ?thesis using assms by (auto simp add: replicate_add) qed qed qed qed lemma step_left_tape_EnlargeBkCtx_eq_C_Bks_Suc: assumes "step0 (s,C@Bk\ z1, r) tm = (s',l',r')" shows "step0 (s,C@Bk\(z1+ Suc z2), r) tm = (s',l'@Bk\Suc z2,r') \ step0 (s,C@Bk\(z1+ Suc z2), r) tm = (s',l'@Bk\z2,r')" proof (cases C) case Nil then have "C = []" . with assms show ?thesis by (metis append.left_neutral step_left_tape_EnlargeBkCtx_eq_Bks) next case (Cons x C') then have "C = x # C'" . then show ?thesis proof (cases x) case Bk then have "x = Bk" . then show ?thesis using assms local.Cons step_left_tape_EnlargeBkCtx_eq_Bk_C_Bks by blast next case Oc then have "x = Oc" . then show ?thesis using assms local.Cons step_left_tape_EnlargeBkCtx_eq_Oc_C_Bks by blast qed qed lemma step_left_tape_EnlargeBkCtx_eq_C_Bks: assumes "step0 (s,C@Bk\ z1, r) tm = (s',l',r')" shows "step0 (s,C@Bk\(z1+ z2), r) tm = (s',l'@Bk\z2,r') \ step0 (s,C@Bk\(z1+ z2), r) tm = (s',l'@Bk\(z2-1),r')" - by (smt step_left_tape_EnlargeBkCtx_eq_C_Bks_Suc One_nat_def Suc_pred add.right_neutral + by (smt (verit) step_left_tape_EnlargeBkCtx_eq_C_Bks_Suc One_nat_def Suc_pred add.right_neutral append.right_neutral assms neq0_conv replicate_empty) (* propagate the step lemmas step_left_tape_EnlargeBkCtx* to steps *) lemma steps_left_tape_EnlargeBkCtx_arbitrary_CL: "steps0 (s, CL @ Bk\z1, r) tm stp = (s', l',r') \ \z3. z3 \ z1 + z2 \ steps0 (s, CL @ Bk\(z1 + z2), r) tm stp = (s',l' @ Bk\z3 ,r')" proof (induct stp arbitrary: s CL z1 r z2 s' l' r') fix s CL z1 r z2 s' l' r' assume "steps0 (s, CL @ Bk \ z1, r) tm 0 = (s', l', r')" show "\z3\z1 + z2. steps0 (s, CL @ Bk \ (z1 + z2), r) tm 0 = (s', l' @ Bk \ z3, r')" by (metis \steps0 (s, CL @ Bk \ z1, r) tm 0 = (s', l', r')\ append.assoc fst_conv le_add2 replicate_add snd_conv steps.simps(1)) next fix stp s CL z1 r z2 s' l' r' assume IV: "\s CL z1 r z2 s' l' r'. steps0 (s, CL @ Bk \ z1, r) tm stp = (s', l', r') \ \z3\z1 + z2. steps0 (s, CL @ Bk \ (z1 + z2), r) tm stp = (s', l' @ Bk \ z3, r')" and minor: "steps0 (s, CL @ Bk \ z1, r) tm (Suc stp) = (s', l', r')" show "\z3\z1 + z2. steps0 (s, CL @ Bk \ (z1 + z2), r) tm (Suc stp) = (s', l' @ Bk \ z3, r')" proof - have F1: "steps0 (s, CL @ Bk \ (z1 + z2), r) tm (Suc stp) = step0 (steps0 (s, CL @ Bk \ (z1 + z2), r) tm stp) tm" by (rule step_red) have "steps0 (s, CL @ Bk \ z1, r) tm (Suc stp) = step0 (steps0 (s, CL @ Bk \ z1, r) tm stp) tm" by (rule step_red) with minor have F3: "step0 (steps0 (s, CL @ Bk \ z1, r) tm stp) tm = (s', l', r')" by auto show "\z3. z3 \ z1 + z2 \ steps0 (s, CL @ Bk \ (z1 + z2), r) tm (Suc stp) = (s', l' @ Bk \ z3, r')" proof (cases "steps0 (s, CL @ Bk \ z1, r) tm stp") case (fields sx lx rx) then have CL: "steps0 (s, CL @ Bk \ z1, r) tm stp = (sx, lx, rx)" . with IV have "\z3'\z1 + z2. steps0 (s, CL @ Bk \ (z1 + z2), r) tm stp = (sx, lx @ Bk \ z3', rx)" by auto then obtain z3' where w_z3': "z3'\z1 + z2 \ steps0 (s, CL @ Bk \ (z1 + z2), r) tm stp = (sx, lx @ Bk \ z3', rx)" by blast moreover have "step0 (sx,lx@Bk\z3', rx) tm = (s',l' @ Bk\(z3') ,r') \ step0 (sx,lx@Bk\z3', rx) tm = (s',l' @ Bk\(z3'-1),r')" proof - from F3 and CL have "step0 (sx, lx@Bk\0, rx) tm = (s', l', r')" by auto then have "step0 (sx,lx@Bk\(0+z3'), rx) tm = (s',l' @ Bk\(z3') ,r') \ step0 (sx,lx@Bk\(0+z3'), rx) tm = (s',l' @ Bk\(z3'-1) ,r')" by (rule step_left_tape_EnlargeBkCtx_eq_C_Bks) (* <-- the step argument *) then show "step0 (sx,lx@Bk\z3', rx) tm = (s',l' @ Bk\(z3') ,r') \ step0 (sx,lx@Bk\z3', rx) tm = (s',l' @ Bk\(z3'-1),r')" by auto qed moreover from w_z3' and F1 have "steps0 (s, CL @ Bk \ (z1 + z2), r) tm (Suc stp) = step0 (sx, lx @ Bk \ z3', rx) tm" by auto ultimately have "steps0 (s, CL @ Bk \ (z1 + z2), r) tm (Suc stp) = (s', l' @ Bk \ z3', r') \ steps0 (s, CL @ Bk \ (z1 + z2), r) tm (Suc stp) = (s', l' @ Bk \ (z3'-1), r')" by auto with w_z3' show ?thesis by (auto simp add: split: if_splits) qed qed qed (* ---------------------------------------------------------------------------------- *) (* The grand lemmas about trailing blanks on the left tape *) (* ---------------------------------------------------------------------------------- *) corollary steps_left_tape_EnlargeBkCtx: " steps0 (s, Bk \ k, r) tm stp = (s', Bk \ l, r') \ \z3. z3 \ k + z2 \ steps0 (s, Bk \ (k + z2), r) tm stp = (s',Bk \ (l + z3), r')" proof - assume "steps0 (s, Bk \ k, r) tm stp = (s', Bk \ l, r')" then have "\z3. z3 \ k + z2 \ steps0 (s, Bk \ (k + z2), r) tm stp = (s',Bk \ l @ Bk \ z3, r')" by (metis append_Nil steps_left_tape_EnlargeBkCtx_arbitrary_CL) then show ?thesis by (simp add: replicate_add) qed corollary steps_left_tape_ShrinkBkCtx_to_NIL: " steps0 (s, Bk \ k, r) tm stp = (s', Bk \ l, r') \ \z3. z3 \ l \ steps0 (s, [], r) tm stp = (s', Bk \ z3, r')" proof - assume A: "steps0 (s, Bk \ k, r) tm stp = (s', Bk \ l, r')" then have F0: "\zb CL'. Bk \ l = CL'@Bk\zb \ steps0 (s,[], r) tm stp = (s',CL',r')" proof (cases k) case 0 then show ?thesis using A append_Nil2 replicate_0 by auto next case (Suc nat) then have "0 < k" by auto with A show ?thesis using steps_left_tape_ShrinkBkCtx_arbitrary_CL by auto qed then obtain zb CL' where w: "Bk \ l = CL'@Bk\zb \ steps0 (s,[], r) tm stp = (s',CL',r')" by blast then show ?thesis by (metis append_same_eq le_add1 length_append length_replicate replicate_add) qed lemma steps_left_tape_Nil_imp_All: "steps0 (s, ([] , r)) p stp = (s', Bk \ k, CR @ Bk \ l) \ \z. \stp k l. (steps0 (s, (Bk\z, r)) p stp) = (s', Bk \ k, CR @ Bk \ l)" proof fix z assume A: "steps0 (s, [], r) p stp = (s', Bk \ k, CR @ Bk \ l)" then have "steps0 (s, Bk \ 0 , r) p stp = (s', Bk \ k, CR @ Bk \ l)" by auto then have "\z3. z3 \ 0 + z \ steps0 (s, Bk \ (0 + z), r) p stp = (s',Bk \ (k + z3), CR @ Bk \ l)" by (rule steps_left_tape_EnlargeBkCtx) then have "\z3. steps0 (s, Bk \ z, r) p stp = (s',Bk \ (k + z3), CR @ Bk \ l)" by auto then obtain z3 where "steps0 (s, Bk \ z, r) p stp = (s',Bk \ (k + z3), CR @ Bk \ l)" by blast then show "\stp k l. steps0 (s, Bk \ z,r) p stp = (s', Bk \ k, CR @ Bk \ l)" by auto qed lemma ex_steps_left_tape_Nil_imp_All: "\stp k l. (steps0 (s, ([] , r)) p stp) = (s', Bk \ k, CR @ Bk \ l) \ \z. \stp k l. (steps0 (s, (Bk\z, r)) p stp) = (s', Bk \ k, CR @ Bk \ l)" proof fix z assume A: "\stp k l. steps0 (s, [], r) p stp = (s', Bk \ k, CR @ Bk \ l)" then obtain stp k l where "steps0 (s, [], r) p stp = (s', Bk \ k, CR @ Bk \ l)" by blast then show "\stp k l. steps0 (s, Bk \ z, r) p stp = (s', Bk \ k, CR @ Bk \ l)" using steps_left_tape_Nil_imp_All by auto qed subsection \Trailing blanks on the right tape do not matter\ (* --------------------------------------------------------------- *) (* Blanks on the right tape, too *) (* --------------------------------------------------------------- *) lemma step_left_tape_Nil_imp_all_trailing_right_Nil: assumes "step0 (s, CL1, [] ) tm = (s', CR1, CR2 )" shows "step0 (s, CL1, [] @ Bk \ y) tm = (s', CR1, CR2 @ Bk \ y) \ step0 (s, CL1, [] @ Bk \ y) tm = (s', CR1, CR2 @ Bk \ (y-1))" proof (cases "fetch tm (s - 0) (read ([] ))") case (Pair a s2) then have A1: "fetch tm (s - 0) (read ([])) = (a, s2)" . show ?thesis proof (cases a) assume "a = WB" from \a = WB\ and assms A1 have "step0 (s, CL1, [] ) tm = (s2, CL1, [Bk])" by auto moreover from \a = WB\ and assms A1 have "step0 (s, CL1, [] @ Bk \ y) tm = (s2, CL1, [Bk]@Bk \ (y-1))" by auto ultimately show ?thesis using assms by auto next assume "a = WO" from \a = WO\ and assms A1 have "step0 (s, CL1, [] ) tm = (s2, CL1, [Oc])" by auto moreover from \a = WO\ and assms A1 have "step0 (s, CL1, [] @ Bk \ y) tm = (s2, CL1, [Oc] @ Bk \ (y-1))" by auto ultimately show ?thesis using assms by auto next assume "a = L" show ?thesis proof (cases CL1) case Nil then have "CL1 = []" . from \CL1 = []\ and \a = L\ and assms A1 have "step0 (s, CL1, [] ) tm = (s2, CL1, [Bk])" by auto moreover from \CL1 = []\ and \a = L\ and assms A1 have "step0 (s, CL1, [] @ Bk \ y) tm = (s2, CL1, Bk # Bk \ y)" by (auto simp add: split: if_splits) ultimately show ?thesis using assms by auto next case (Cons c cs) then have "CL1 = c # cs" . from \CL1 = c # cs \ and \a = L\ and assms A1 have "step0 (s, CL1, [] ) tm = (s2, cs, [c])" by auto moreover from \CL1 = c # cs\ and \a = L\ and assms A1 have "step0 (s, CL1, [] @ Bk \ y) tm = (s2, cs, c # Bk \ y)" by (auto simp add: split: if_splits) ultimately show ?thesis using assms by auto qed next assume "a = Nop" from \a = Nop\ and assms and A1 have "step0 (s, CL1, [] ) tm = (s2, CL1, [] )" by auto moreover from \a = Nop\ and assms and A1 have "step0 (s, CL1, [] @ Bk \ y) tm = (s2, CL1, [] @ Bk \ y)" by auto ultimately show ?thesis using assms by auto next assume "a = R" from \a = R\ and assms A1 have "step0 (s, CL1, [] ) tm = (s2, Bk#CL1, [] )" by auto moreover from \a = R\ and assms A1 have "step0 (s, CL1, [] @ Bk \ y) tm = (s2, Bk#CL1, []@ Bk \ (y-1) )" by auto ultimately show ?thesis using assms by auto qed qed lemma step_left_tape_Nil_imp_all_trailing_right_Cons: assumes "step0 (s, CL1, rx#rs ) tm = (s', CR1, CR2 )" shows "step0 (s, CL1, rx#rs @ Bk \ y) tm = (s', CR1, CR2 @ Bk \ y)" proof (cases "fetch tm (s - 0) (read (rx#rs ))") case (Pair a s2) then have A1: "fetch tm (s - 0) (read (rx#rs )) = (a, s2)" . show ?thesis proof (cases a) assume "a = WB" from \a = WB\ and assms A1 have "step0 (s, CL1, rx#rs ) tm = (s2, CL1, Bk#rs )" by auto moreover from \a = WB\ and assms A1 have "step0 (s, CL1, rx#rs @ Bk \ y) tm = (s2, CL1, Bk#rs @ Bk \ y)" by auto ultimately show ?thesis using assms by auto next assume "a = WO" from \a = WO\ and assms A1 have "step0 (s, CL1, rx#rs ) tm = (s2, CL1, Oc#rs )" by auto moreover from \a = WO\ and assms A1 have "step0 (s, CL1, rx#rs @ Bk \ y) tm = (s2, CL1, Oc#rs @ Bk \ y)" by auto ultimately show ?thesis using assms by auto next assume "a = L" show ?thesis proof (cases CL1) case Nil then have "CL1 = []" . show ?thesis proof - from \a = L\ and \CL1 = []\ and assms A1 have "step0 (s, CL1, rx#rs ) tm = (s2, CL1, Bk#rx#rs )" by auto moreover from \a = L\ and \CL1 = []\ and assms A1 have "step0 (s, CL1, rx#rs @ Bk \ y) tm = (s2, CL1, Bk#rx#rs @ Bk \ y)" by auto ultimately show ?thesis using assms by auto qed next case (Cons c cs) then have "CL1 = c # cs" . show ?thesis proof - from \a = L\ and \CL1 = c # cs\ and assms A1 have "step0 (s, CL1, rx#rs ) tm = (s2, cs, c#rx#rs )" by auto moreover from \a = L\ and \CL1 = c # cs\ and assms A1 have "step0 (s, CL1, rx#rs @ Bk \ y) tm = (s2, cs, c#rx#rs @ Bk \ y)" by auto ultimately show ?thesis using assms by auto qed qed next assume "a = R" from \a = R\ and assms A1 have "step0 (s, CL1, rx#rs ) tm = (s2, rx#CL1, rs )" by auto moreover from \a = R\ and assms A1 have "step0 (s, CL1, rx#rs @ Bk \ y) tm = (s2, rx#CL1, rs @ Bk \ y)" by auto ultimately show ?thesis using assms by auto next assume "a = Nop" from \a = Nop\ and assms A1 have "step0 (s, CL1, rx#rs ) tm = (s2, CL1, rx#rs )" by (auto simp add: split: if_splits) moreover from \a = Nop\ and assms A1 have "step0 (s, CL1, rx#rs @ Bk \ y) tm = (s2, CL1, rx#rs @ Bk \ y)" by auto ultimately show ?thesis using assms by auto qed qed lemma step_left_tape_Nil_imp_all_trailing_right: assumes "step0 (s, CL1, r ) tm = (s', CR1, CR2 )" shows "step0 (s, CL1, r @ Bk \ y) tm = (s', CR1, CR2 @ Bk \ y) \ step0 (s, CL1, r @ Bk \ y) tm = (s', CR1, CR2 @ Bk \ (y-1))" proof (cases r) case Nil then show ?thesis using step_left_tape_Nil_imp_all_trailing_right_Nil assms by auto next case (Cons a list) then show ?thesis using step_left_tape_Nil_imp_all_trailing_right_Cons assms by auto qed (* propagate to steps *) lemma steps_left_tape_Nil_imp_all_trailing_right: "steps0 (s, CL1, r ) tm stp = (s', CR1, CR2) \ \x1 x2. y = x1 + x2 \ steps0 (s, CL1, r @ Bk \ y ) tm stp = (s', CR1, CR2 @ Bk \ x2)" proof (induct stp arbitrary: s CL1 r y s' CR1 CR2) case 0 then show ?case by (simp add: "0.prems" steps_left_tape_Nil_imp_All) next fix stp s CL1 r y s' CR1 CR2 assume IV: "\s CL1 r y s' CR1 CR2. steps0 (s, CL1, r) tm stp = (s', CR1, CR2) \ \x1 x2. y = x1 + x2 \ steps0 (s, CL1, r @ Bk \ y) tm stp = (s', CR1, CR2 @ Bk \ x2)" and major: "steps0 (s, CL1, r) tm (Suc stp) = (s', CR1, CR2)" show "\x1 x2. y = x1 + x2 \ steps0 (s, CL1, r @ Bk \ y) tm (Suc stp) = (s', CR1, CR2 @ Bk \ x2)" proof - have F1: "steps0 (s, CL1, r @ Bk \ y) tm (Suc stp) = step0 (steps0 (s, CL1, r @ Bk \ y) tm stp) tm" by (rule step_red) have "steps0 (s, CL1, r) tm (Suc stp) = step0 (steps0 (s, CL1, r) tm stp) tm" by (rule step_red) with major have F3: "step0 (steps0 (s, CL1, r) tm stp) tm = (s', CR1, CR2)" by auto show "\x1 x2. y = x1 + x2 \ steps0 (s, CL1, r @ Bk \ y) tm (Suc stp) = (s', CR1, CR2 @ Bk \ x2)" proof (cases y) case 0 then have "y = 0" . with major show ?thesis by auto next case (Suc y') then have "y = Suc y'" . then show ?thesis proof (cases "steps0 (s, CL1, r) tm stp") case (fields sx lx rx) then have C: "steps0 (s, CL1, r) tm stp = (sx, lx, rx)" . with major and IV have F0: "\x1' x2'. y = x1' + x2' \ steps0 (s, CL1, r @ Bk \ y) tm stp = (sx, lx, rx @ Bk \ x2')" by auto then obtain x1' x2' where w_x1'_x2': " y = x1' + x2' \ steps0 (s, CL1, r @ Bk \ y) tm stp = (sx, lx, rx @ Bk \ x2')" by blast moreover have "step0 (sx, lx, rx @ Bk \ x2') tm = (s', CR1, CR2 @ Bk \ x2') \ step0 (sx, lx, rx @ Bk \ x2') tm = (s', CR1, CR2 @ Bk \ (x2'-1))" proof - from F3 and C have F5: "step0 (sx, lx, rx) tm = (s', CR1, CR2)" by auto then have "step0 (sx, lx, rx @ Bk \ x2') tm = (s', CR1, CR2 @ Bk \ x2') \ step0 (sx, lx, rx @ Bk \ x2') tm = (s', CR1, CR2 @ Bk \ (x2'-1))" by (rule step_left_tape_Nil_imp_all_trailing_right) then show "step0 (sx, lx, rx @ Bk \ x2') tm = (s', CR1, CR2 @ Bk \ x2') \ step0 (sx, lx, rx @ Bk \ x2') tm = (s', CR1, CR2 @ Bk \ (x2'-1))" by auto qed moreover from w_x1'_x2' and F1 have "steps0 (s, CL1, r @ Bk \ y) tm (Suc stp) = step0 (sx, lx, rx @ Bk \ x2') tm" by auto ultimately have F5: "y = x1' + x2' \ (steps0 (s, CL1, r @ Bk \ y) tm (Suc stp) = (s', CR1, CR2 @ Bk \ x2') \ steps0 (s, CL1, r @ Bk \ y) tm (Suc stp) = (s', CR1, CR2 @ Bk \ (x2' - 1)))" by auto with w_x1'_x2' show ?thesis proof (cases x2') case 0 with F5 have "y = x1' + 0 \ steps0 (s, CL1, r @ Bk \ y) tm (Suc stp) = (s', CR1, CR2 @ Bk \ 0)" by (auto simp add: split: if_splits) with w_x1'_x2' show ?thesis by (auto simp add: split: if_splits) next case (Suc x2'') with w_x1'_x2' show ?thesis using F5 add_Suc_right add_Suc_shift diff_Suc_1 by fastforce qed qed qed qed qed (* ---------------------------------------------------------------------------------- *) (* The grand lemmas about trailing blanks on the right tape *) (* ---------------------------------------------------------------------------------- *) lemma ex_steps_left_tape_Nil_imp_All_left_and_right: "(\kr lr. steps0 (1, ([] , r )) p stp = (0, Bk \ kr, r' @ Bk \ lr)) \ \kl ll. \kr lr. steps0 (1, (Bk \ kl, r @ Bk \ ll)) p stp = (0, Bk \ kr, r' @ Bk \ lr)" proof - assume "(\kr lr. steps0 (1, ([] , r )) p stp = (0, Bk \ kr, r' @ Bk \ lr))" then obtain kr lr where w_kr_lr: "steps0 (1, ([] , r )) p stp = (0, Bk \ kr, r' @ Bk \ lr)" by blast then have "steps0 (1, (Bk \ 0 , r )) p stp = (0, Bk \ kr, r' @ Bk \ lr)" by auto then have "\kl. \z3. z3 \ 0 + kl \ steps0 (1, Bk \ (0 + kl), r) p stp = (0,Bk \ (kr + z3), r' @ Bk \ lr)" using steps_left_tape_EnlargeBkCtx using plus_nat.add_0 w_kr_lr by blast then have "\kl. \z3. z3 \ kl \ steps0 (1, Bk \ kl, r) p stp = (0,Bk \ (kr + z3), r' @ Bk \ lr)" by auto then have F0: "\kl. \z4. steps0 (1, Bk \ kl, r ) p stp = (0,Bk \ z4 , r' @ Bk \ lr)" by auto show ?thesis proof fix kl show " \ll. \kr lr. steps0 (1, Bk \ kl, r @ Bk \ ll) p stp = (0, Bk \ kr, r' @ Bk \ lr)" proof fix ll show "\kr lr. steps0 (1, Bk \ kl, r @ Bk \ ll) p stp = (0, Bk \ kr, r' @ Bk \ lr)" proof - from F0 have "\z4. steps0 (1, Bk \ kl, r ) p stp = (0,Bk \ z4 , r' @ Bk \ lr)" by auto then obtain z4 where w_z4: "steps0 (1, Bk \ kl, r ) p stp = (0,Bk \ z4 , r' @ Bk \ lr)" by blast then have "\x1 x2. ll = x1 + x2 \ steps0 (1, Bk \ kl , r @ Bk \ ll ) p stp = (0, Bk \ z4 , r' @ Bk \ lr @ Bk \ x2)" using steps_left_tape_Nil_imp_all_trailing_right using append_assoc by fastforce then show "\kr lr. steps0 (1, Bk \ kl, r @ Bk \ ll) p stp = (0, Bk \ kr, r' @ Bk \ lr)" by (metis replicate_add) qed qed qed qed end diff --git a/thys/Universal_Turing_Machine/HaltingProblems_K_H.thy b/thys/Universal_Turing_Machine/HaltingProblems_K_H.thy --- a/thys/Universal_Turing_Machine/HaltingProblems_K_H.thy +++ b/thys/Universal_Turing_Machine/HaltingProblems_K_H.thy @@ -1,567 +1,567 @@ (* Title: thys/HaltingProblems_K_H.thy Author: Franz Regensburger (FABR) 02/2022 *) subsection \Undecidability of Halting Problems\ theory HaltingProblems_K_H imports SimpleGoedelEncoding SemiIdTM TuringReducible begin (* -------------------------------------------------------------------------- *) (* declare adjust.simps[simp del] declare seq_tm.simps [simp del] declare shift.simps[simp del] declare composable_tm.simps[simp del] declare step.simps[simp del] declare steps.simps[simp del] *) subsubsection \A locale for variations of the Halting Problem\ text \ The following locale assumes that there is an injective coding function \t2c\ from Turing machines to natural numbers. In this locale, we will show that the Special Halting Problem K1 and the General Halting Problem H1 are not Turing decidable. \ locale hpk = fixes t2c :: "tprog0 \ nat" assumes t2c_inj: "inj t2c" begin text \The function @{term tm_to_nat} is a witness that the locale hpk is inhabited.\ interpretation tm_to_nat: hpk "tm_to_nat :: tprog0 \ nat" proof unfold_locales show "inj tm_to_nat" by (rule inj_tm_to_nat) qed text \We define the function @{term c2t} as the unique inverse of the injective function @{term t2c}. \ (* Note 1: * It does not matter how we define c2t n if \(\p. t2c p = n). * We never need to know that value * since we are only interested in the equation * * "c2t (t2c p) = p" *) (* Note 1: * We use Hilbert' non-constructive operators for * definite choice \ (iota) (THE in Isabelle/HOL) * indefinite choice \ (epsilon) (SOME in Isabelle/HOL) *) definition c2t :: "nat \ instr list" where "c2t = (\n. if (\p. t2c p = n) then (THE p. t2c p = n) else (SOME p. True) )" lemma t2c_inj': "inj_on t2c {x. True}" by (auto simp add: t2c_inj ) lemma c2t_comp_t2c_eq: "c2t (t2c p) = p" proof - have "\p\{x. True}. c2t (t2c p) = p" proof (rule encode_decode_A_eq[OF t2c_inj']) show "c2t = (\w. if \t\{x. True}. t2c t = w then THE t. t \ {x. True} \ t2c t = w else SOME t. t \ {x. True})" by (auto simp add: c2t_def) qed then show ?thesis by auto qed subsubsection \Undecidability of the Special Halting Problem K1\ (* Just as a reminder: definition TMC_has_num_res :: "tprog0 \ nat list \ bool" where "TMC_has_num_res p ns \ \ \tap. tap = ([], ) \ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" lemma TMC_yields_num_res_unfolded_into_Hoare_halt: "TMC_yields_num_res tm ns n \ \(\tap. tap = ([], ))\ tm \(\tap. \k l. tap = (Bk \ k, @ Bk\ l))\" *) definition K1 :: "(nat list) set" where "K1 \ {nl. (\n. nl = [n] \ TMC_has_num_res (c2t n) [n]) }" text \Assuming the existence of a Turing Machine K1D1, which is able to decide the set K1, we derive a contradiction using the machine @{term "tm_semi_id_eq0"}. Thus, we show that the {\em Special Halting Problem K1} is not Turing decidable. \label{sec_K1_H1} The proof uses a diagonal argument. \ (* some convenience lemmas will keep the main proof quit neat *) lemma mk_composable_decider_K1D1: assumes "\K1D1'. (\nl. (nl \ K1 \TMC_yields_num_res K1D1' nl (1::nat)) \ (nl \ K1 \TMC_yields_num_res K1D1' nl (0::nat) ))" shows "\K1D1'. (\nl. composable_tm0 K1D1' \ (nl \ K1 \TMC_yields_num_res K1D1' nl (1::nat)) \ (nl \ K1 \TMC_yields_num_res K1D1' nl (0::nat) ))" proof - from assms have "\K1D1'. (\nl. (nl \ K1 \\\tap. tap = ([], )\ K1D1' \\tap. \k l. tap = (Bk \ k, [Oc,Oc] @ Bk \ l)\) \ (nl \ K1 \\\tap. tap = ([], )\ K1D1' \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\) )" unfolding TMC_yields_num_res_unfolded_into_Hoare_halt by (simp add: tape_of_nat_def) (* create a composable version of the potentially non-composable machine K1D1' *) then obtain K1D1' where "(\nl. (nl \ K1 \ \\tap. tap = ([], )\ K1D1' \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\) \ (nl \ K1 \ \\tap. tap = ([], )\ K1D1' \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\))" by blast then have "composable_tm0 (mk_composable0 K1D1') \ (\nl. (nl \ K1 \ \\tap. tap = ([], )\ mk_composable0 K1D1' \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\) \ (nl \ K1 \ \\tap. tap = ([], )\ mk_composable0 K1D1' \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\))" using Hoare_halt_tm_impl_Hoare_halt_mk_composable0 composable_tm0_mk_composable0 by blast then have "composable_tm0 (mk_composable0 K1D1') \ (\nl. (nl \ K1 \ TMC_yields_num_res (mk_composable0 K1D1') nl (1::nat) ) \ (nl \ K1 \ TMC_yields_num_res (mk_composable0 K1D1') nl (0::nat) ))" unfolding TMC_yields_num_res_unfolded_into_Hoare_halt by (simp add: tape_of_nat_def) then show ?thesis by auto qed lemma res_1_fed_into_tm_semi_id_eq0_loops: assumes "composable_tm0 D" and "TMC_yields_num_res D nl (1::nat)" shows "TMC_loops (D |+| tm_semi_id_eq0) nl" unfolding TMC_loops_def proof fix stp show "\ is_final (steps0 (1, [], ) (D |+| tm_semi_id_eq0) stp)" using assms tm_semi_id_eq0_loops'' Hoare_plus_unhalt Hoare_unhalt_def tape_of_nat_def tape_of_list_def TMC_yields_num_res_unfolded_into_Hoare_halt by simp qed lemma loops_imp_has_no_res: "TMC_loops tm [n] \ \ TMC_has_num_res tm [n]" proof - assume "TMC_loops tm [n]" then show "\ TMC_has_num_res tm [n]" using TMC_has_num_res_iff TMC_loops_def by blast qed lemma yields_res_imp_has_res: "TMC_yields_num_res tm [n] (m::nat) \ TMC_has_num_res tm [n]" proof - assume "TMC_yields_num_res tm [n] (m::nat)" then show "TMC_has_num_res tm [n]" by (metis TMC_has_num_res_iff TMC_yields_num_res_def is_finalI) qed lemma res_0_fed_into_tm_semi_id_eq0_yields_0: assumes "composable_tm0 D" and "TMC_yields_num_res D nl (0::nat)" shows "TMC_yields_num_res (D |+| tm_semi_id_eq0) nl 0" unfolding TMC_yields_num_res_unfolded_into_Hoare_halt using assms Hoare_plus_halt tm_semi_id_eq0_halts'' tape_of_nat_def tape_of_list_def TMC_yields_num_res_unfolded_into_Hoare_halt by simp (* here is the main lemma about the Halting problem K1 *) lemma existence_of_decider_K1D1_for_K1_imp_False: assumes major: "\K1D1'. (\nl. (nl \ K1 \TMC_yields_num_res K1D1' nl (1::nat)) \ (nl \ K1 \TMC_yields_num_res K1D1' nl (0::nat) ))" shows False proof - (* first, create a composable version of the potentially non-composable machine K1D1' *) from major have "\K1D1'. (\nl. composable_tm0 K1D1' \ (nl \ K1 \ TMC_yields_num_res K1D1' nl (1::nat)) \ (nl \ K1 \ TMC_yields_num_res K1D1' nl (0::nat) ))" by (rule mk_composable_decider_K1D1) then obtain K1D1 where w_K1D1: "\nl. composable_tm0 K1D1 \ (nl \ K1 \ TMC_yields_num_res K1D1 nl (1::nat)) \ (nl \ K1 \ TMC_yields_num_res K1D1 nl (0::nat) )" by blast (* second, using our composable decider K1D1 we construct a machine tm_contra that will contradict our major assumption using a diagonal argument *) define tm_contra where "tm_contra = K1D1 |+| tm_semi_id_eq0" (* third, we derive the crucial lemma "c2t (t2c tm_contra) = tm_contra" *) have c2t_comp_t2c_eq_for_tm_contra: "c2t (t2c tm_contra) = tm_contra" by (auto simp add: c2t_comp_t2c_eq) (* now, we derive the contradiction *) show False (* using the classical axiom TND: A \ \A, * we proof by cases: [t2c tm_contra] \ K1 \ [t2c tm_contra] \ K1 *) proof (cases "[t2c tm_contra] \ K1") case True from \[t2c tm_contra] \ K1\ and w_K1D1 have "TMC_yields_num_res K1D1 [t2c tm_contra] (1::nat)" by auto then have "TMC_loops tm_contra [t2c tm_contra]" using res_1_fed_into_tm_semi_id_eq0_loops w_K1D1 tm_contra_def by blast then have "\(TMC_has_num_res tm_contra [t2c tm_contra])" using loops_imp_has_no_res by blast then have "\(TMC_has_num_res (c2t (t2c tm_contra))) [t2c tm_contra]" by (auto simp add: c2t_comp_t2c_eq_for_tm_contra) then have "[t2c tm_contra] \ K1" by (auto simp add: K1_def) with \[t2c tm_contra] \ K1\ show False by auto next case False from \[t2c tm_contra] \ K1\ and w_K1D1 have "TMC_yields_num_res K1D1 [t2c tm_contra] (0::nat)" by auto then have "TMC_yields_num_res tm_contra [t2c tm_contra] (0::nat)" using res_0_fed_into_tm_semi_id_eq0_yields_0 w_K1D1 tm_contra_def by auto then have "TMC_has_num_res tm_contra [t2c tm_contra]" using yields_res_imp_has_res by blast then have "TMC_has_num_res (c2t (t2c tm_contra)) [t2c tm_contra]" by (auto simp add: c2t_comp_t2c_eq_for_tm_contra) then have "[t2c tm_contra] \ K1" by (auto simp add: K1_def) with \[t2c tm_contra] \ K1\ show False by auto qed qed (* FABR NOTE: We are not able to show a result like "\(turing_decidable H2)" Reason: Our notion of Turing decidability is based on: turing_decidable :: "(nat list) set \ bool" The lemma of theory TuringUnComputable_H2 lemma existence_of_decider_H2D0_for_H2_imp_False: assumes "\H2D0'. (\nl (tm::instr list). ((tm,nl) \ H2 \\\tap. tap = ([], <(code tm, nl)>)\ H2D0' \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk\l)\) \ ((tm,nl) \ H2 \\\tap. tap = ([], <(code tm, nl)>)\ H2D0' \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @Bk\l)\) )" shows False is about pairs (tm,nl) :: ((instr list) \ (nat list)) set definition H2 :: "((instr list) \ (nat list)) set" where "H2 \ {(tm,nl). TMC_has_num_res tm nl }" We need to derive a variation of this result where we use a set of lists [code tm, nl] :: (nat list) set that fits into our setting. definition H1 :: "(nat list) set" where "H1 \ {nl. (\n m. nl = [n,m] \ TMC_has_num_res (c2t n) [m]) }" *) subsubsection \The Special Halting Problem K1 is reducible to the General Halting Problem H1\ text \The proof is by reduction of \K1\ to \H1\. \ (* Convenience lemmas for the reduction of K1 to H1 *) definition H1 :: "(nat list) set" where "H1 \ {nl. (\n m. nl = [n,m] \ TMC_has_num_res (c2t n) [m]) }" lemma NilNotIn_K1: "[] \ K1" unfolding K1_def using CollectD list.simps(3) by auto lemma NilNotIn_H1: "[] \ H1" unfolding H1_def using CollectD list.simps(3) by auto lemma tm_strong_copy_total_correctness_Nil': "length nl = 0 \ TMC_yields_num_list_res tm_strong_copy nl []" unfolding TMC_yields_num_list_res_unfolded_into_Hoare_halt proof - assume "length nl = 0" then have "\\tap. tap = ([], ) \ tm_strong_copy \\tap. tap = ([Bk,Bk,Bk,Bk],[]) \" using tm_strong_copy_total_correctness_Nil by simp then have F1: "\\tap. tap = ([], )\ tm_strong_copy \\tap. tap = (Bk \ 4, <[]> @ Bk \ 0)\" by (metis One_nat_def Suc_1 \length nl = 0\ append_Nil length_0_conv numeral_4_eq_4 numeral_eqs_upto_12(2) replicate_0 replicate_Suc tape_of_list_empty) show "\\tap. tap = ([], )\ tm_strong_copy \\tap. \k l. tap = (Bk \ k, <[]::nat list> @ Bk \ l)\" proof (rule Hoare_haltI) fix l::"cell list" fix r::"cell list" assume "(l, r) = ([], )" show "\n. is_final (steps0 (1, l, r) tm_strong_copy n) \ (\tap. \k l. tap = (Bk \ k, <[]::nat list> @ Bk \ l)) holds_for steps0 (1, l, r) tm_strong_copy n" - by (smt F1 Hoare_haltE \(l, r) = ([], )\ holds_for.elims(2) holds_for.simps) + by (smt (verit) F1 Hoare_haltE \(l, r) = ([], )\ holds_for.elims(2) holds_for.simps) qed qed lemma tm_strong_copy_total_correctness_len_eq_1': "length nl = 1 \ TMC_yields_num_list_res tm_strong_copy nl [hd nl, hd nl]" unfolding TMC_yields_num_list_res_unfolded_into_Hoare_halt proof - assume "length nl = 1" then show "\\tap. tap = ([], ) \ tm_strong_copy \ \tap. \k l. tap = (Bk \ k, <[hd nl, hd nl]> @ Bk \ l) \" using tm_strong_copy_total_correctness_len_eq_1 by simp qed lemma tm_strong_copy_total_correctness_len_gt_1': "1 < length nl \ TMC_yields_num_list_res tm_strong_copy nl [hd nl]" unfolding TMC_yields_num_list_res_unfolded_into_Hoare_halt proof - assume "1 < length nl" then have "\\tap. tap = ([], ) \ tm_strong_copy \ \tap. \l. tap = ([Bk,Bk], <[hd nl]> @ Bk \ l) \" using tm_strong_copy_total_correctness_len_gt_1 by simp then show "\\tap. tap = ([], ) \ tm_strong_copy \ \tap. \k l. tap = (Bk \ k, <[hd nl]> @ Bk \ l) \" - by (smt Hoare_haltE Hoare_haltI One_nat_def Pair_inject Pair_inject holds_for.elims(2) + by (smt (verit, ccfv_threshold) Hoare_haltE Hoare_haltI One_nat_def Pair_inject Pair_inject holds_for.elims(2) holds_for.simps is_final.elims(2) replicate.simps(1) replicate.simps(2)) qed (* --- *) theorem turing_reducible_K1_H1: "turing_reducible K1 H1" unfolding turing_reducible_unfolded_into_TMC_yields_condition proof - have "\nl::nat list. \ml::nat list. TMC_yields_num_list_res tm_strong_copy nl ml \ (nl \ K1 \ ml \ H1)" proof fix nl::"nat list" have "length nl = 0 \ length nl = 1 \ 1 < length nl" by arith then show "\ml. TMC_yields_num_list_res tm_strong_copy nl ml \ (nl \ K1) = (ml \ H1)" proof assume "length nl = 0" then have "TMC_yields_num_list_res tm_strong_copy nl []" by (rule tm_strong_copy_total_correctness_Nil') moreover have "(nl \ K1) = ([] \ H1)" using NilNotIn_H1 NilNotIn_K1 \length nl = 0\ length_0_conv by blast ultimately show "\ml. TMC_yields_num_list_res tm_strong_copy nl ml \ (nl \ K1) = (ml \ H1)" by blast next assume "length nl = 1 \ 1 < length nl" then show "\ml. TMC_yields_num_list_res tm_strong_copy nl ml \ (nl \ K1) = (ml \ H1)" proof assume "1 < length nl" then have "TMC_yields_num_list_res tm_strong_copy nl [hd nl]" by (rule tm_strong_copy_total_correctness_len_gt_1') moreover have "(nl \ K1) = ([hd nl] \ H1)" using H1_def K1_def \1 < length nl\ by auto ultimately show "\ml. TMC_yields_num_list_res tm_strong_copy nl ml \ (nl \ K1) = (ml \ H1)" by blast next assume "length nl = 1" then have "TMC_yields_num_list_res tm_strong_copy nl [hd nl, hd nl]" by (rule tm_strong_copy_total_correctness_len_eq_1') moreover have "(nl \ K1) = ([hd nl, hd nl] \ H1)" - by (smt CollectD Cons_eq_append_conv H1_def K1_def One_nat_def \length nl = 1\ + by (smt (verit) CollectD Cons_eq_append_conv H1_def K1_def One_nat_def \length nl = 1\ append_Cons diff_Suc_1 hd_Cons_tl length_0_conv length_tl list.inject mem_Collect_eq not_Cons_self2 self_append_conv2 zero_neq_one) ultimately show "\ml. TMC_yields_num_list_res tm_strong_copy nl ml \ (nl \ K1) = (ml \ H1)" by blast qed qed qed then show "\tm. \nl. \ml. TMC_yields_num_list_res tm nl ml \ (nl \ K1) = (ml \ H1)" by auto qed (* --------------------------------------------------------------------------------------------- *) subsubsection \Corollaries about the undecidable sets K1 and H1\ corollary not_Turing_decidable_K1: "\(turing_decidable K1)" proof assume "turing_decidable K1" then have "(\D. (\nl. (nl \ K1 \TMC_yields_num_res D nl (1::nat)) \ (nl \ K1 \TMC_yields_num_res D nl (0::nat) )))" by (auto simp add: turing_decidable_unfolded_into_TMC_yields_conditions tape_of_nat_def) with existence_of_decider_K1D1_for_K1_imp_False show False by blast qed corollary not_Turing_decidable_H1: "\turing_decidable H1" proof (rule turing_reducible_AB_and_non_decA_imp_non_decB) show "turing_reducible K1 H1" by (rule turing_reducible_K1_H1) next show "\ turing_decidable K1" by (rule not_Turing_decidable_K1) qed (* --------------------------------------------------------------------------------------------- *) subsubsection \Proof variant: The special Halting Problem K1 is not Turing Decidable\ text\Assuming the existence of a Turing Machine K1D0, which is able to decide the set K1, we derive a contradiction using the machine @{term "tm_semi_id_gt0"}. Thus, we show that the {\em Special Halting Problem K1} is not Turing decidable. The proof uses a diagonal argument. \label{sec_K1_v}\ lemma existence_of_decider_K1D0_for_K1_imp_False: assumes "\K1D0'. (\nl. (nl \ K1 \TMC_yields_num_res K1D0' nl (0::nat)) \ (nl \ K1 \TMC_yields_num_res K1D0' nl (1::nat) ))" shows False proof - from assms have "\K1D0'. (\nl. (nl \ K1 \\\tap. tap = ([], )\ K1D0' \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\) \ (nl \ K1 \\\tap. tap = ([], )\ K1D0' \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\) )" unfolding TMC_yields_num_res_unfolded_into_Hoare_halt by (simp add: tape_of_nat_def) then obtain K1D0' where "(\nl. (nl \ K1 \ \\tap. tap = ([], )\ K1D0' \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\) \ (nl \ K1 \ \\tap. tap = ([], )\ K1D0' \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\))" by blast (* first, create a composable version of the potentially non-composable machine K1D0' *) then have "composable_tm0 (mk_composable0 K1D0') \ (\nl. (nl \ K1 \ \\tap. tap = ([], )\ mk_composable0 K1D0' \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\) \ (nl \ K1 \ \\tap. tap = ([], )\ mk_composable0 K1D0' \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\))" using Hoare_halt_tm_impl_Hoare_halt_mk_composable0 composable_tm0_mk_composable0 by blast then have "\K1D0. composable_tm0 K1D0 \ (\nl. (nl \ K1 \ \\tap. tap = ([], )\ K1D0 \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\) \ (nl \ K1 \ \\tap. tap = ([], )\ K1D0 \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\))" by blast then obtain K1D0 where w_K1D0: "composable_tm0 K1D0 \ (\nl. (nl \ K1 \ \\tap. tap = ([], )\ K1D0 \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\) \ (nl \ K1 \ \\tap. tap = ([], )\ K1D0 \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\))" by blast define tm_contra where "tm_contra = K1D0 |+| tm_semi_id_gt0" (* second, we derive some theorems from our assumptions *) then have c2t_comp_t2c_TM_eq_for_tm_contra: "c2t (t2c tm_contra) = tm_contra" by (auto simp add: c2t_comp_t2c_eq) (* now, we derive the contradiction *) show False proof (cases "[t2c tm_contra] \ K1") case True from \[t2c tm_contra] \ K1\ and w_K1D0 have "\\tap. tap = ([], <[t2c tm_contra]>)\ K1D0 \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\" by auto then have "\\tap. tap = ([], <[t2c tm_contra]>)\ tm_contra \" unfolding tm_contra_def proof (rule Hoare_plus_unhalt) show "\\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\ tm_semi_id_gt0 \" by (rule tm_semi_id_gt0_loops'') next from w_K1D0 show "composable_tm0 K1D0" by auto qed then have "\\tap. tap = ([], <[t2c tm_contra]>)\ c2t (t2c tm_contra) \" by (auto simp add: c2t_comp_t2c_TM_eq_for_tm_contra) then have "\\\tap. tap = ([], <[t2c tm_contra]>)\ c2t (t2c tm_contra) \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" using Hoare_halt_impl_not_Hoare_unhalt by blast then have "\( TMC_has_num_res (c2t (t2c tm_contra)) [t2c tm_contra])" by (auto simp add: TMC_has_num_res_def) then have "[t2c tm_contra] \ K1" by (auto simp add: K1_def) with \[t2c tm_contra] \ K1\ show False by auto next case False from \[t2c tm_contra] \ K1\ and w_K1D0 have "\\tap. tap = ([], <[t2c tm_contra]>)\ K1D0 \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\" by auto then have "\\tap. tap = ([], <[t2c tm_contra]>)\ tm_contra \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\" unfolding tm_contra_def proof (rule Hoare_plus_halt) show "\\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\ tm_semi_id_gt0 \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\" by (rule tm_semi_id_gt0_halts'') next from w_K1D0 show "composable_tm0 K1D0" by auto qed then have "\\tap. tap = ([], <[t2c tm_contra]>)\ c2t (t2c tm_contra) \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\" by (auto simp add: c2t_comp_t2c_TM_eq_for_tm_contra) then have " TMC_has_num_res (c2t (t2c tm_contra)) [t2c tm_contra]" by (auto simp add: Hoare_halt_with_OcOc_imp_std_tap tape_of_nat_def) then have "[t2c tm_contra] \ K1" by (auto simp add: K1_def) with \[t2c tm_contra] \ K1\ show False by auto qed qed end (* locale hpk *) end diff --git a/thys/Universal_Turing_Machine/Numerals.thy b/thys/Universal_Turing_Machine/Numerals.thy --- a/thys/Universal_Turing_Machine/Numerals.thy +++ b/thys/Universal_Turing_Machine/Numerals.thy @@ -1,874 +1,874 @@ (* Title: thys/Numerals.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten Further contributions by Franz Regensburger (FABR) 02/2022 : - Re-ordering of sections *) section \Encoding of Natural Numbers\ theory Numerals imports ComposedTMs BlanksDoNotMatter begin subsection \A class for generating numerals\ class tape = fixes tape_of :: "'a \ cell list" ("<_>" 100) instantiation nat::tape begin definition tape_of_nat where "tape_of_nat (n::nat) \ Oc \ (Suc n)" instance by standard end type_synonym nat_list = "nat list" instantiation list::(tape) tape begin (* fun tape_of_nat_list :: "nat list \ cell list" The type ('a::tape) list is needed for lemma tape_of_list_empty in Abacus_Mopup *) fun tape_of_nat_list :: "('a::tape) list \ cell list" where "tape_of_nat_list [] = []" | "tape_of_nat_list [n] = " | "tape_of_nat_list (n#ns) = @ Bk # (tape_of_nat_list ns)" definition tape_of_list where "tape_of_list \ tape_of_nat_list" instance by standard end instantiation prod:: (tape, tape) tape begin fun tape_of_nat_prod :: "('a::tape) \ ('b::tape) \ cell list" where "tape_of_nat_prod (n, m) = @ [Bk] @ " definition tape_of_prod where "tape_of_prod \ tape_of_nat_prod" instance by standard end subsection \Some lemmas about numerals used for rewriting\ (* Lemma tape_of_list_empty is needed in Abacus_Mopup.thy *) lemma tape_of_list_empty[simp]: "<[]> = ([]::cell list)" by (simp add: tape_of_list_def) lemma tape_of_nat_list_cases2: "<(nl::nat list)> = [] \ (\r'. = Oc # r')" by (induct rule: tape_of_nat_list.induct)(auto simp add: tape_of_nat_def tape_of_list_def) subsection \Unique decomposition of standard tapes\ text \Some lemmas about unique decomposition of tapes in standard halting configuration. \ lemma OcSuc_lemma: "Oc # Oc \ n1 = Oc \ n2 \ Suc n1 = n2" proof (induct n1 arbitrary: n2) case 0 then have A: "Oc # Oc \ 0 = Oc \ n2" . then show ?case proof - from A have "[Oc] = Oc \ n2" by auto moreover have "[Oc] = Oc \ n2 \ Suc 0 = n2" by (induct n2)auto ultimately show ?case by auto qed next fix n1 n2 assume IV1: "\n2. Oc # Oc \ n1 = Oc \ n2 \ Suc n1 = n2" and IV2: "Oc # Oc \ Suc n1 = Oc \ n2" show "Suc (Suc n1) = n2" proof (cases n2) case 0 then have "Oc \ n2 = []" by auto with IV2 have False by auto then show ?thesis by auto next case (Suc n3) then have "n2 = Suc n3" . with IV2 have "Oc # Oc \ Suc n1 = Oc \ (Suc n3)" by auto then have "Oc # Oc # Oc \ n1 = Oc # Oc \ n3" by auto then have "Oc # Oc \ n1 = Oc \ n3" by auto with IV1 have "Suc n1 = n3" by auto then have "Suc (Suc n1) = Suc n3" by auto with \n2 = Suc n3\ show ?thesis by auto qed qed lemma inj_tape_of_list: "() = () \ n1 = n2" by (induct n1 arbitrary: n2) (auto simp add: OcSuc_lemma tape_of_nat_def) lemma inj_repl_Bk: "Bk \ k1 = Bk \ k2 \ k1 = k2" by auto lemma last_of_numeral_is_Oc: "last () = Oc" by (auto simp add: tape_of_nat_def) lemma hd_of_numeral_is_Oc: "hd () = Oc" by (auto simp add: tape_of_nat_def) lemma rev_replicate: "rev (Bk \ l1) = (Bk \ l1)" by auto lemma rev_numeral: "rev () = " by (induct n)(auto simp add: tape_of_nat_def) lemma drop_Bk_prefix: "n < l \ hd (drop n ((Bk \ l) @ xs)) = Bk" by (induct n arbitrary: l xs)(auto) lemma unique_Bk_postfix: " @ Bk \ l1 = @ Bk \ l2 \ l1 = l2" proof - assume " @ Bk \ l1 = @ Bk \ l2" then have "rev ( @ Bk \ l1) = rev ( @ Bk \ l2)" by auto then have "rev (Bk \ l1) @ rev () = rev (Bk \ l2) @ rev ()" by auto then have A: "(Bk \ l1) @ () = (Bk \ l2) @ ()" by (auto simp add: rev_replicate rev_numeral) then show "l1 = l2" proof (cases "l1 = l2") case True then show ?thesis by auto next case False then have "l1 < l2 \ l2 < l1" by auto then show ?thesis proof assume "l1 < l2 " then have False proof - have "hd (drop l1 (Bk \ l1) @ ()) = hd ( )" by auto also have "... = Oc" by (auto simp add: hd_of_numeral_is_Oc) finally have F1: "hd (drop l1 (Bk \ l1) @ ()) = Oc" . from \l1 < l2\have "hd (drop l1 ((Bk \ l2) @ ())) = Bk" by (auto simp add: drop_Bk_prefix) with A have "hd (drop l1 ((Bk \ l1) @ ())) = Bk" by auto with F1 show False by auto qed then show ?thesis by auto next assume "l2 < l1" then have False proof - have "hd (drop l2 (Bk \ l2) @ ()) = hd ( )" by auto also have "... = Oc" by (auto simp add: hd_of_numeral_is_Oc) finally have F2: "hd (drop l2 (Bk \ l2) @ ()) = Oc" . from \l2 < l1\ have "hd (drop l2 ((Bk \ l1) @ ())) = Bk" by (auto simp add: drop_Bk_prefix) with A have "hd (drop l2 ((Bk \ l2) @ ())) = Bk" by auto with F2 show False by auto qed then show ?thesis by auto qed qed qed lemma unique_decomp_tap: assumes "(lx, @ Bk \ l1) = (ly, @ Bk \ l2)" shows "lx=ly \ n1=n2 \ l1=l2" proof from assms show "lx = ly " by auto next show "n1 = n2 \ l1 = l2" proof from assms have major: " @ Bk \ l1 = @ Bk \ l2" by auto then have "rev ( @ Bk \ l1) = rev ( @ Bk \ l2)" by auto then have "rev (Bk \ l1) @ rev () = rev (Bk \ l2) @ rev ()" by auto then have A: "(Bk \ l1) @ () = (Bk \ l2) @ ()" by (auto simp add: rev_replicate rev_numeral) then show "n1 = n2" proof - from major have "l1 = l2" by (rule unique_Bk_postfix) with A have "() = ()" by auto then show "n1 = n2" by (rule inj_tape_of_list) qed next from assms have major: " @ Bk \ l1 = @ Bk \ l2" by auto then show "l1 = l2" by (rule unique_Bk_postfix) qed qed lemma unique_decomp_std_tap: assumes "(Bk \ k1, @ Bk \ l1) = (Bk \ k2, @ Bk \ l2)" shows "k1=k2 \ n1=n2 \ l1=l2" proof from assms have "Bk \ k1 = Bk \ k2" by auto then show "k1 = k2" by auto next show "n1 = n2 \ l1 = l2" proof from assms have major: " @ Bk \ l1 = @ Bk \ l2" by auto then have "rev ( @ Bk \ l1) = rev ( @ Bk \ l2)" by auto then have "rev (Bk \ l1) @ rev () = rev (Bk \ l2) @ rev ()" by auto then have A: "(Bk \ l1) @ () = (Bk \ l2) @ ()" by (auto simp add: rev_replicate rev_numeral) then show "n1 = n2" proof - from major have "l1 = l2" by (rule unique_Bk_postfix) with A have "() = ()" by auto then show "n1 = n2" by (rule inj_tape_of_list) qed next from assms have major: " @ Bk \ l1 = @ Bk \ l2" by auto then show "l1 = l2" by (rule unique_Bk_postfix) qed qed subsection \Lists of numerals never contain two consecutive blanks\ definition noDblBk:: "cell list \ bool" where "noDblBk cs \ \i. Suc i < length cs \ cs!i = Bk \ cs!(Suc i) = Oc" lemma noDblBk_Bk_Oc_rep: "noDblBk (Oc \ n1)" by (simp add: noDblBk_def ) lemma noDblBk_Bk_imp_Oc: "\noDblBk cs; Suc i < length cs; cs!i = Bk \ \ cs!(Suc i) = Oc" by (auto simp add: noDblBk_def) lemma noDblBk_imp_noDblBk_Oc_cons: "noDblBk cs \ noDblBk (Oc#cs)" - by (smt Suc_less_eq Suc_pred add.right_neutral add_Suc_right cell.exhaust list.size(4) + by (smt (verit) Suc_less_eq Suc_pred add.right_neutral add_Suc_right cell.exhaust list.size(4) neq0_conv noDblBk_Bk_imp_Oc noDblBk_def nth_Cons_0 nth_Cons_Suc) lemma noDblBk_Numeral: "noDblBk ()" by (auto simp add: noDblBk_def tape_of_nat_def) lemma noDblBk_Nil: "noDblBk []" by (auto simp add: noDblBk_def) lemma noDblBk_Singleton: "noDblBk (<[n::nat]>)" by (auto simp add: noDblBk_def tape_of_nat_def tape_of_list_def) (* the next two lemmas are for the inductive step in lemma noDblBk_tape_of_nat_list *) lemma tape_of_nat_list_cons_eq:"nl \ [] \ <(a::nat) # nl> = @ Bk # " by (metis list.exhaust tape_of_list_def tape_of_nat_list.simps(3)) lemma noDblBk_cons_cons: "noDblBk(<(x::nat) # xs>) \ noDblBk( @ Bk # )" proof - assume F0: "noDblBk ()" have F1: "hd() = Oc" by (metis hd_append hd_of_numeral_is_Oc list.sel(1) tape_of_nat_list_cons_eq tape_of_list_def tape_of_nat_list.simps(2) tape_of_nat_list_cases2) have F2: " = Oc \ (Suc a)" by (auto simp add: tape_of_nat_def) have "noDblBk ()" by (auto simp add: noDblBk_Numeral) with F0 and F1 and F2 show ?thesis unfolding noDblBk_def proof - assume A1: "\i. Suc i < length () \ ! i = Bk \ ! Suc i = Oc" and A2: "hd () = Oc" and A3: " = Oc \ Suc a" and A4: "\i. Suc i < length () \ ! i = Bk \ ! Suc i = Oc" show "\i. Suc i < length ( @ Bk # ) \ ( @ Bk # ) ! i = Bk \ ( @ Bk # ) ! Suc i = Oc" proof fix i show "Suc i < length ( @ Bk # ) \ ( @ Bk # ) ! i = Bk \ ( @ Bk # ) ! Suc i = Oc" proof assume "Suc i < length ( @ Bk # ) \ ( @ Bk # ) ! i = Bk" then show "( @ Bk # ) ! Suc i = Oc" proof assume A5: "Suc i < length ( @ Bk # )" and A6: "( @ Bk # ) ! i = Bk" show "( @ Bk # ) ! Suc i = Oc" proof - from A5 have "Suc i < length () \ Suc i = length () \ Suc i = Suc (length ()) \ (Suc (length ()) < Suc i \ Suc i < length ( @ Bk # ))" by auto then show ?thesis proof assume "Suc i < length ()" with A1 A2 A3 A4 A5 show ?thesis by (simp add: nth_append') next assume "Suc i = length () \ Suc i = Suc (length ()) \ Suc (length ()) < Suc i \ Suc i < length ( @ Bk # )" then show "( @ Bk # ) ! Suc i = Oc" proof assume "Suc i = length ()" with A1 A2 A3 A4 A5 A6 show "( @ Bk # ) ! Suc i = Oc" by (metis lessI nth_Cons_Suc nth_append' nth_append_length replicate_append_same) next assume "Suc i = Suc (length ()) \ Suc (length ()) < Suc i \ Suc i < length ( @ Bk # )" then show "( @ Bk # ) ! Suc i = Oc" proof assume "Suc i = Suc (length ())" with A1 A2 A3 A4 A5 A6 show "( @ Bk # ) ! Suc i = Oc" by (metis One_nat_def Suc_eq_plus1 length_Cons length_append list.collapse list.size(3) nat_neq_iff nth_Cons_0 nth_Cons_Suc nth_append_length_plus) next assume A7: "Suc (length ()) < Suc i \ Suc i < length ( @ Bk # )" have F3: "( @ Bk # ) = (( @ [Bk]) @ )" by auto have F4: "\n. (( @ [Bk]) @ )! (length ( @ [Bk]) + n) = ()!n" using nth_append_length_plus by blast from A7 have "\m. i = Suc (length ()) + m" by arith then obtain m where w_m: "i = Suc (length ()) + m" by blast with A7 have F5: "Suc m < length ()" by auto from w_m F4 have F6: "(( @ [Bk]) @ ) ! i = ()! m" by auto with F5 A7 have F7: "(( @ [Bk]) @ ) ! (Suc i) = ()! (Suc m)" by (metis F4 add_Suc_right length_append_singleton w_m) from A6 and F6 have "()! m = Bk" by auto with A1 and F5 have "()! (Suc m) = Oc" by auto with F7 have "(( @ [Bk]) @ ) ! (Suc i) = Oc" by auto with F3 show "( @ Bk # )! (Suc i) = Oc" by auto qed qed qed qed qed qed qed qed qed theorem noDblBk_tape_of_nat_list: "noDblBk()" proof (induct nl) case Nil then show ?case by (auto simp add: noDblBk_def tape_of_nat_def tape_of_list_def) next case (Cons a nl) then have IV: "noDblBk ()" . show "noDblBk ()" proof (cases nl) case Nil then show ?thesis by (auto simp add: noDblBk_def tape_of_nat_def tape_of_list_def) next case (Cons x xs) then have "nl = x # xs" . show ?thesis proof - from \nl = x # xs\ have "noDblBk () = noDblBk( @ Bk # )" by (auto simp add: tape_of_nat_list_cons_eq) also with IV and \nl = x # xs\ have "... = True" using noDblBk_cons_cons by auto finally show ?thesis by auto qed qed qed lemma hasDblBk_L1: "\ CR = rs @ [Bk] @ Bk # rs'; noDblBk CR \ \ False" by (metis add_diff_cancel_left' append.simps(2) append_assoc cell.simps(2) length_Cons length_append length_append_singleton noDblBk_def nth_append_length zero_less_Suc zero_less_diff) lemma hasDblBk_L2: "\ C = Bk # cls; noDblBk C \ \ cls = [] \ (\cls'. cls = Oc#cls')" by (metis (full_types) append_Cons append_Nil cell.exhaust hasDblBk_L1 neq_Nil_conv) lemma hasDblBk_L3: "\ noDblBk C ; C = C1 @ (Bk#C2) \ \ C2 = [] \ (\C3. C2 = Oc#C3)" by (metis (full_types) append_Cons append_Nil cell.exhaust hasDblBk_L1 neq_Nil_conv) lemma hasDblBk_L4: assumes "noDblBk CL" and "r = Bk # rs" and "r = rev ls1 @ Oc # rss" and "CL = ls1 @ ls2" shows "ls2 = [] \ (\bs. ls2 = Oc#bs)" proof - from \r = Bk # rs\ have "last ls1 = Bk" proof (cases ls1) case Nil then have "ls1=[]" . with \r = rev ls1 @ Oc # rss\ have "r = Oc # rss" by auto with \r = Bk # rs\ have False by auto then show ?thesis by auto next case (Cons b bs) then have "ls1 = b#bs" . with \r = rev ls1 @ Oc # rss\ and \r = Bk # rs\ show ?thesis by (metis \r = rev ls1 @ Oc # rss\ last_appendR last_snoc list.simps(3) rev.simps(2) rev_append rev_rev_ident) qed show ?thesis proof (cases ls2) case Nil then show ?thesis by auto next case (Cons b bs) then have "ls2 = b # bs" . show ?thesis proof (cases b) case Bk then have "b = Bk" . with \ls2 = b # bs\ have "ls2 = Bk # bs" by auto with \CL = ls1 @ ls2\ have "CL = ls1 @ Bk # bs" by auto then have "CL = butlast ls1 @ [Bk] @ Bk # ls2" by (metis \last ls1 = Bk\ \noDblBk CL\ \r = Bk # rs\ \r = rev ls1 @ Oc # rss\ append_butlast_last_id append_eq_append_conv2 append_self_conv2 cell.distinct(1) hasDblBk_L1 list.inject rev_is_Nil_conv) with \noDblBk CL\ have False using hasDblBk_L1 by blast then show ?thesis by auto next case Oc with \ls2 = b # bs\ show ?thesis by auto qed qed qed lemma hasDblBk_L5: assumes "noDblBk CL" and "r = Bk # rs" and "r = rev ls1 @ Oc # rss" and "CL = ls1 @ [Bk]" shows False using assms hasDblBk_L4 by blast lemma noDblBk_cases: assumes "noDblBk C" and "C = C1 @ C2" and "C2 = [] \ P" and "C2 = [Bk] \ P" and "\C3. C2 = Bk#Oc#C3 \ P" and "\C3. C2 = Oc#C3 \ P" shows "P" proof - have "C2 = [] \ C2 = [Bk] \ (\C3. C2 = Bk#Oc#C3 \ C2 = Bk#Bk#C3 \ C2 = Oc#C3)" by (metis (full_types) cell.exhaust list.exhaust) then show ?thesis using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) hasDblBk_L3 list.distinct(1) by blast qed subsection \Unique decomposition of tapes containing lists of numerals\ text \A lemma about appending lists of numerals.\ lemma append_numeral_list: "\ (nl1::nat list) \ []; nl2 \ [] \ \ = @[Bk]@" proof (induct nl1 arbitrary: nl2) case Nil then show ?case by blast next fix a::nat fix nl1::"nat list" fix nl2::"nat list" assume IH: "\nl2. \nl1 \ []; nl2 \ []\ \ = @ [Bk] @ " and minor1: "a # nl1 \ []" and minor2: "nl2 \ []" show "<(a # nl1) @ nl2> = @ [Bk] @ " proof (cases nl1) assume "nl1 = []" then show "<(a # nl1) @ nl2> = @ [Bk] @ " by (metis append_Cons append_Nil minor2 tape_of_list_def tape_of_nat_list.simps(2) tape_of_nat_list_cons_eq) next fix na nl1s assume "nl1 = na # nl1s" then have "nl1 \ []" by auto have "<(a # nl1) @ nl2> = " by auto also with \nl1 \ []\ and \nl1 = na # nl1s\ have "... = @[Bk]@<(nl1 @ nl2)>" by (simp add: tape_of_nat_list_cons_eq) also with \nl1 \ []\ and minor2 and IH have "... = @[Bk]@ @ [Bk] @ " by auto finally have "<(a # nl1) @ nl2> = @[Bk]@ @ [Bk] @ " by auto moreover with \nl1 \ []\ and minor2 have " @ [Bk] @ = @[Bk]@ @ [Bk] @ " by (simp add: tape_of_nat_list_cons_eq) ultimately show "<(a # nl1) @ nl2> = @ [Bk] @ " by auto qed qed text \A lemma about reverting lists of numerals.\ lemma rev_numeral_list: "rev() = <(rev nl)>" proof (induct nl) case Nil then show ?case by (simp) next fix a::nat fix nl::"nat list" assume IH: "rev () = " show "rev () = " proof (rule tape_of_nat_list.cases[of nl]) assume "nl = []" then have "rev () = rev ()" by (simp add: tape_of_list_def ) with \nl = []\ show ?thesis by (simp add: rev_numeral tape_of_list_def ) next fix n assume "nl = [n]" then have "rev () = rev ()" by auto also have "... = rev (@[Bk]@)" by (simp add: \nl = [n]\ tape_of_list_def tape_of_nat_list_cons_eq) also have "... = rev()@[Bk]@rev()" by simp also with IH have "... = @[Bk]@rev()" by auto also with IH \nl = [n]\ have "... = " by (simp add: Cons_eq_append_conv hd_rev rev_numeral tape_of_list_def ) finally show "rev () = " by auto next fix n v va assume "nl = n # v # va" then have "rev () = rev ()" by auto also with \nl = n # v # va\ have "... = rev()@[Bk]@rev()" by (simp add: tape_of_list_def ) also with IH have "... = @[Bk]@rev()" by auto finally have "rev () = @[Bk]@rev()" by auto moreover have " = @[Bk]@rev()" proof - have " = " by auto also from \nl = n # v # va\ have " = @[Bk]@rev()" by (metis list.simps(3) append_numeral_list rev_is_Nil_conv rev_numeral tape_of_list_def tape_of_nat_list.simps(2)) finally show " = @[Bk]@rev()" by auto qed ultimately show ?thesis by auto qed qed text \Some more lemmas about unique decomposition of tapes that contain lists of numerals.\ lemma unique_Bk_postfix_numeral_list_Nil: "<[]> @ Bk \ l1 = @ Bk \ l2 \ [] = yl" proof (induct yl arbitrary: l1 l2) case Nil then show ?case by auto next fix a fix yl:: "nat list" fix l1 l2 assume IV: "\l1 l2. <[]> @ Bk \ l1 = @ Bk \ l2 \ [] = yl" and major: "<[]> @ Bk \ l1 = @ Bk \ l2" then show "[] = a # yl" by (metis append.assoc append.simps(1) append.simps(2) cell.distinct(1) list.sel(1) list.simps(3) replicate_Suc replicate_app_Cons_same tape_of_list_def tape_of_nat_def tape_of_nat_list.elims tape_of_nat_list.simps(3) tape_of_nat_list_cases2) qed lemma nonempty_list_of_numerals_neq_BKs: " \ Bk \ l" by (metis append_Nil append_Nil2 list.simps(3) replicate_0 tape_of_list_def tape_of_list_empty unique_Bk_postfix_numeral_list_Nil) (* ENHANCE: simplify proof: use cases "a < b \ a = b \ b < a" only once (earlier in the proof) *) lemma unique_Bk_postfix_nonempty_numeral_list: "\ xl \ []; yl \ []; @ Bk \ l1 = @ Bk \ l2 \ \ xl = yl" proof (induct xl arbitrary: l1 yl l2) fix l1 yl l2 assume "<[]> @ Bk \ l1 = @ Bk \ l2" then show "[] = yl" using unique_Bk_postfix_numeral_list_Nil by auto next fix a:: nat fix xl:: "nat list" fix l1 fix yl:: "nat list" fix l2 assume IV: "\l1 yl l2. \xl \ []; yl \ []; @ Bk \ l1 = @ Bk \ l2\ \ xl = yl" and minor_xl: "a # xl \ []" and minor_yl: "yl \ []" and major: " @ Bk \ l1 = @ Bk \ l2" show "a # xl = yl" proof (cases yl) case Nil then show ?thesis by (metis major tape_of_list_empty unique_Bk_postfix_numeral_list_Nil) next fix b:: nat fix ys:: "nat list" assume Ayl: "yl = b # ys" have "a # xl = b # ys" proof show "a = b \ xl = ys" proof (cases xl) case Nil then have "xl = []" . from major and \yl = b # ys\ have " @ Bk \ l1 = @ Bk \ l2" by auto with minor_xl and \xl = []\ have " = " by (simp add: local.Nil tape_of_list_def) with major and Ayl have " @ Bk \ l1 = @ Bk \ l2" by auto show "a = b \ xl = ys" proof (cases ys) case Nil then have "ys = []" . then have " = " by (simp add: local.Nil tape_of_list_def) with \ @ Bk \ l1 = @ Bk \ l2\ have " @ Bk \ l1 = @ Bk \ l2" by auto then have "a = b" by (metis append_same_eq inj_tape_of_list unique_Bk_postfix) with \xl = []\ and \ys = []\ show ?thesis by auto next fix c fix ys':: "nat list" assume "ys = c# ys'" then have " = @ Bk # " by (simp add: Ayl tape_of_list_def) with \ @ Bk \ l1 = @ Bk \ l2\ have " @ Bk \ l1 = @ Bk # @ Bk \ l2" by simp show "a = b \ xl = ys" proof (cases l1) case 0 then have "l1 = 0" . with \ @ Bk \ l1 = @ Bk # @ Bk \ l2\ have " = @ Bk # @ Bk \ l2" by auto then have False by (metis "0" \ @ Bk \ l1 = @ Bk # @ Bk \ l2\ cell.distinct(1) length_Cons length_append length_replicate less_add_same_cancel1 nth_append_length nth_replicate tape_of_nat_def zero_less_Suc) then show ?thesis by auto next case (Suc l1') then have "l1 = Suc l1'" . then have " @ Bk \ l1 = @ (Bk #Bk \ l1')" by simp then have F1: "( @ Bk \ l1)!(Suc a) = Bk" by (metis cell.distinct(1) cell.exhaust length_replicate nth_append_length tape_of_nat_def) have "a < b \ a = b \ b < a" by arith (* move upwards in the proof: do cases only once *) then show "a = b \ xl = ys" proof assume "a < b" with \ @ Bk \ l1 = @ Bk # @ Bk \ l2\ have "( @ Bk # @ Bk \ l2)!(Suc a) \ Bk" by (simp add: hd_of_numeral_is_Oc nth_append' tape_of_nat_def) with F1 and \ @ Bk \ l1 = @ Bk # @ Bk \ l2\ have False by (simp add: \( @ Bk \ l1) ! Suc a = Bk\) then show "a = b \ xl = ys" by auto next assume "a = b \ b < a" then show ?thesis proof assume "b < a" with \ @ Bk \ l1 = @ Bk # @ Bk \ l2\ have "( @ Bk \ l1)!(Suc a) \ Bk" by (metis Suc_mono cell.distinct(1) length_replicate nth_append' nth_append_length nth_replicate tape_of_nat_def) with F1 have False by auto then show ?thesis by auto next assume "a=b" then have False using \xl = []\ and \ys = c# ys'\ and \ @ Bk \ l1 = @ Bk # @ Bk \ l2\ using \ @ Bk \ l1 = @ Bk # Bk \ l1'\ list.distinct(1) list.inject same_append_eq self_append_conv2 tape_of_list_empty unique_Bk_postfix_numeral_list_Nil by fastforce then show ?thesis by auto qed qed qed qed next fix a'::nat fix xs:: "nat list" assume "xl = a'#xs" from major and \yl = b # ys\ have " @ Bk \ l1 = @ Bk \ l2" by auto from \xl = a'#xs\ have " = @ Bk # " by (simp add: tape_of_list_def) with \ @ Bk \ l1 = @ Bk \ l2\ have "( @ Bk # ) @ Bk \ l1 = @ Bk \ l2" by auto then have F2: " @ [Bk] @ ( @ Bk \ l1) = @ Bk \ l2" by auto then have F3: "( @ [Bk] @ ( @ Bk \ l1))!(Suc a) = Bk" by (metis Ayl append.simps(1) append.simps(2) length_replicate nth_append_length tape_of_nat_def) show "a = b \ xl = ys" proof (cases ys) case Nil then have "ys = []" . then have " = " by (simp add: local.Nil tape_of_list_def) with F2 have " @ [Bk] @ ( @ Bk \ l1) = @ Bk \ l2" by auto have "a < b \ a = b \ b < a" by arith (* move upwards in the proof: do cases only once *) then have False proof assume "a < b" then have "( @ Bk \ l2)!(Suc a) \ Bk" by (simp add: \a < b\ nth_append' tape_of_nat_def) with F2 and F3 show False using \ @ [Bk] @ @ Bk \ l1 = @ Bk \ l2\ by auto next assume "a = b \ b < a" then show False proof assume "b < a" show False proof (cases l2) case 0 then have "l2 = 0" . then have " @ Bk \ l2 = " by auto with \ @ [Bk] @ ( @ Bk \ l1) = @ Bk \ l2\ have " @ [Bk] @ ( @ Bk \ l1) = " by auto with \xl = a'#xs\ and \b < a\ have "length () < length ( @ [Bk] @ ( @ Bk \ l1))" by (metis inj_tape_of_list le_add1 length_append less_irrefl nat_less_le nth_append' nth_equalityI) with \b < a\ have "b < length ( @ [Bk] @ ( @ Bk \ l1))" by (simp add: \ @ [Bk] @ @ Bk \ l1 = \ tape_of_nat_def) with \ @ [Bk] @ ( @ Bk \ l1) = \ show False using Suc_less_SucD \b < a\ \length () < length ( @ [Bk] @ @ Bk \ l1)\ length_replicate not_less_iff_gr_or_eq tape_of_nat_def by auto next fix l2' assume "l2 = Suc l2'" with \ @ [Bk] @ ( @ Bk \ l1) = @ Bk \ l2\ have " @ [Bk] @ ( @ Bk \ l1) = @ [Bk] @ Bk \ l2'" by auto from \b < a\ have "( @ [Bk] @ Bk \ l2')!(Suc b) = Bk" by (metis append_Cons length_replicate nth_append_length tape_of_nat_def) moreover from \b < a\ have "( @ [Bk] @ ( @ Bk \ l1))!(Suc b) \ Bk" by (simp add: Suc_mono last_of_numeral_is_Oc nth_append' tape_of_nat_def) ultimately show False using \ @ [Bk] @ @ Bk \ l1 = @ [Bk] @ Bk \ l2'\ by auto qed next assume "a=b" with \ @ [Bk] @ ( @ Bk \ l1) = @ Bk \ l2\ and \xl = a'#xs\ show False by (metis append_Nil append_eq_append_conv2 list.sel(3) list.simps(3) local.Nil same_append_eq tape_of_list_empty tl_append2 tl_replicate unique_Bk_postfix_numeral_list_Nil) qed qed then show ?thesis by auto next fix c fix ys':: "nat list" assume "ys = c# ys'" from \ys = c# ys'\ have "() = ( @ [Bk] @ )" by (simp add: Ayl tape_of_list_def ) with F2 have " @ [Bk] @ ( @ Bk \ l1) = @ [Bk] @ @ Bk \ l2" by auto have "a < b \ a = b \ b < a" by arith (* move upwards in the proof: do cases only once *) then show "a = b \ xl = ys" proof assume "a < b" with \a < b\ have "( @ [Bk] @ @ Bk \ l2) !(Suc a) \ Bk" by (simp add: hd_of_numeral_is_Oc nth_append' tape_of_nat_def) with F3 and \ @ [Bk] @ ( @ Bk \ l1) = @ [Bk] @ @ Bk \ l2\ have False by auto then show ?thesis by auto next assume "a = b \ b < a" then show ?thesis proof assume "b < a" from \b < a\ and \ @ [Bk] @ ( @ Bk \ l1) = @ [Bk] @ @ Bk \ l2\ have "( @ [Bk] @ @ Bk \ l2) ! Suc b = Bk" by (metis append_Cons cell.distinct(1) cell.exhaust length_replicate nth_append_length tape_of_nat_def) moreover from \b < a\ and \ @ [Bk] @ ( @ Bk \ l1) = @ [Bk] @ @ Bk \ l2\ have "( @ [Bk] @ ( @ Bk \ l1)) ! Suc b \ Bk" by (metis Suc_mono cell.distinct(1) length_replicate nth_append' nth_replicate tape_of_nat_def) ultimately have False using \ @ [Bk] @ ( @ Bk \ l1) = @ [Bk] @ @ Bk \ l2\ by auto then show ?thesis by auto next assume "a = b" with \ @ [Bk] @ ( @ Bk \ l1) = @ [Bk] @ @ Bk \ l2\ have " @ Bk \ l1 = @ Bk \ l2" by auto moreover from \xl = a'#xs\ have "xl \ []" by auto moreover from \ys = c# ys'\ have "ys \ []" by auto ultimately have "xl = ys" using IV by auto with \a = b\ show ?thesis by auto qed qed qed qed qed with \yl = b # ys\ show ?thesis by auto qed qed corollary unique_Bk_postfix_numeral_list: " @ Bk \ l1 = @ Bk \ l2 \ xl = yl" by (metis append_Nil tape_of_list_def tape_of_list_empty unique_Bk_postfix_nonempty_numeral_list unique_Bk_postfix_numeral_list_Nil) text \Some more lemmas about noDblBks in lists of numerals.\ lemma numeral_list_head_is_Oc: "(nl::nat list) \ [] \ hd () = Oc" proof - assume A: "(nl::nat list) \ []" then have "(\r'. = Oc # r')" using append_Nil tape_of_list_empty tape_of_nat_list_cases2 unique_Bk_postfix_numeral_list_Nil by fastforce then obtain r' where w_r': " = Oc # r'" by blast then show "hd () = Oc" by auto qed lemma numeral_list_last_is_Oc: "(nl::nat list) \ [] \ last () = Oc" proof - assume A: "(nl::nat list) \ []" then have " = " by auto also have "... = rev ()" by (auto simp add: rev_numeral_list) finally have " = rev ()" by auto moreover from A have "hd () = Oc" by (auto simp add: numeral_list_head_is_Oc) with A have "last (rev ()) = Oc" by (simp add: last_rev) ultimately show ?thesis by (simp add: \last (rev ()) = Oc\) qed lemma noDblBk_tape_of_nat_list_imp_noDblBk_tl: "noDblBk () \ noDblBk (tl ())" proof (cases "") case Nil then show ?thesis by (simp add: local.Nil noDblBk_Nil) next fix a nls assume "noDblBk ()" and " = a # nls" then have "noDblBk (a # nls)" by auto then have "\i. Suc i < length (a # nls) \ (a # nls) ! i = Bk \ (a # nls) ! Suc i = Oc" using noDblBk_def by auto then have "noDblBk (nls)" using noDblBk_def by auto with \ = a # nls\ show "noDblBk (tl ())" using noDblBk_def by auto qed lemma noDblBk_tape_of_nat_list_cons_imp_noDblBk_tl: "noDblBk (a # ) \ noDblBk ()" proof - assume "noDblBk (a # )" then have "\i. Suc i < length (a # ) \ (a # ) ! i = Bk \ (a # ) ! Suc i = Oc" using noDblBk_def by auto then show "noDblBk ()" using noDblBk_def by auto qed lemma noDblBk_tape_of_nat_list_imp_noDblBk_cons_Bk: "(nl::nat list) \ [] \ noDblBk ([Bk] @ )" proof - assume "(nl::nat list) \ []" then have major: "<(0::nat) # nl> = <0::nat> @ Bk # " using tape_of_nat_list_cons_eq by auto moreover have "noDblBk (<(0::nat) # nl>)" by (rule noDblBk_tape_of_nat_list) ultimately have "noDblBk (<0::nat> @ Bk # )" by auto then have "noDblBk (Oc # Bk # )" by (simp add: cell.exhaust noDblBk_tape_of_nat_list tape_of_nat_def) then show ?thesis using major by (metis append_eq_Cons_conv empty_replicate list.sel(3) noDblBk_tape_of_nat_list_imp_noDblBk_tl replicate_Suc self_append_conv2 tape_of_nat_def ) qed end diff --git a/thys/Universal_Turing_Machine/Recursive.thy b/thys/Universal_Turing_Machine/Recursive.thy --- a/thys/Universal_Turing_Machine/Recursive.thy +++ b/thys/Universal_Turing_Machine/Recursive.thy @@ -1,2869 +1,2869 @@ (* Title: thys/Recursive.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten Modifications: Franz Regensburger (FABR) 08/2022 Added LaTeX sections and comments *) section \Compilation of Recursive Functions into Abacus Programs\ theory Recursive imports Abacus Rec_Def Abacus_Hoare begin fun addition :: "nat \ nat \ nat \ abc_prog" where "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, Inc m, Goto 4]" fun mv_box :: "nat \ nat \ abc_prog" where "mv_box m n = [Dec m 3, Inc n, Goto 0]" text \The compilation of \z\-operator.\ definition rec_ci_z :: "abc_inst list" where "rec_ci_z \ [Goto 1]" text \The compilation of \s\-operator.\ definition rec_ci_s :: "abc_inst list" where "rec_ci_s \ (addition 0 1 2 [+] [Inc 1])" text \The compilation of \id i j\-operator\ fun rec_ci_id :: "nat \ nat \ abc_inst list" where "rec_ci_id i j = addition j i (i + 1)" fun mv_boxes :: "nat \ nat \ nat \ abc_inst list" where "mv_boxes ab bb 0 = []" | "mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] mv_box (ab + n) (bb + n)" fun empty_boxes :: "nat \ abc_inst list" where "empty_boxes 0 = []" | "empty_boxes (Suc n) = empty_boxes n [+] [Dec n 2, Goto 0]" fun cn_merge_gs :: "(abc_inst list \ nat \ nat) list \ nat \ abc_inst list" where "cn_merge_gs [] p = []" | "cn_merge_gs (g # gs) p = (let (gprog, gpara, gn) = g in gprog [+] mv_box gpara p [+] cn_merge_gs gs (Suc p))" subsection \Definition of the compiler rec\_ci\ text \ The compiler of recursive functions, where \rec_ci recf\ return \(ap, arity, fp)\, where \ap\ is the Abacus program, \arity\ is the arity of the recursive function \recf\, \fp\ is the amount of memory which is going to be used by \ap\ for its execution. \ fun rec_ci :: "recf \ abc_inst list \ nat \ nat" where "rec_ci z = (rec_ci_z, 1, 2)" | "rec_ci s = (rec_ci_s, 1, 3)" | "rec_ci (id m n) = (rec_ci_id m n, m, m + 2)" | "rec_ci (Cn n f gs) = (let cied_gs = map (\ g. rec_ci g) gs in let (fprog, fpara, fn) = rec_ci f in let pstr = Max (set (Suc n # fn # (map (\ (aprog, p, n). n) cied_gs))) in let qstr = pstr + Suc (length gs) in (cn_merge_gs cied_gs pstr [+] mv_boxes 0 qstr n [+] mv_boxes pstr 0 (length gs) [+] fprog [+] mv_box fpara pstr [+] empty_boxes (length gs) [+] mv_box pstr n [+] mv_boxes qstr 0 n, n, qstr + n))" | "rec_ci (Pr n f g) = (let (fprog, fpara, fn) = rec_ci f in let (gprog, gpara, gn) = rec_ci g in let p = Max (set ([n + 3, fn, gn])) in let e = length gprog + 7 in (mv_box n p [+] fprog [+] mv_box n (Suc n) [+] (([Dec p e] [+] gprog [+] [Inc n, Dec (Suc n) 3, Goto 1]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gprog + 4)]), Suc n, p + 1))" | "rec_ci (Mn n f) = (let (fprog, fpara, fn) = rec_ci f in let len = length (fprog) in (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3), Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn))" subsection \Correctness of the compiler rec\_ci\ declare rec_ci.simps [simp del] rec_ci_s_def[simp del] rec_ci_z_def[simp del] rec_ci_id.simps[simp del] mv_boxes.simps[simp del] mv_box.simps[simp del] addition.simps[simp del] declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] abc_step_l.simps[simp del] inductive_cases terminate_pr_reverse: "terminate (Pr n f g) xs" inductive_cases terminate_z_reverse[elim!]: "terminate z xs" inductive_cases terminate_s_reverse[elim!]: "terminate s xs" inductive_cases terminate_id_reverse[elim!]: "terminate (id m n) xs" inductive_cases terminate_cn_reverse[elim!]: "terminate (Cn n f gs) xs" inductive_cases terminate_mn_reverse[elim!]:"terminate (Mn n f) xs" fun addition_inv :: "nat \ nat list \ nat \ nat \ nat \ nat list \ bool" where "addition_inv (as, lm') m n p lm = (let sn = lm ! n in let sm = lm ! m in lm ! p = 0 \ (if as = 0 then \ x. x \ lm ! m \ lm' = lm[m := x, n := (sn + sm - x), p := (sm - x)] else if as = 1 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm - x - 1), p := (sm - x - 1)] else if as = 2 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm - x), p := (sm - x - 1)] else if as = 3 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm - x), p := (sm - x)] else if as = 4 then \ x. x \ lm ! m \ lm' = lm[m := x, n := (sn + sm), p := (sm - x)] else if as = 5 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm), p := (sm - x - 1)] else if as = 6 then \ x. x < lm ! m \ lm' = lm[m := Suc x, n := (sn + sm), p := (sm - x - 1)] else if as = 7 then lm' = lm[m := sm, n := (sn + sm)] else False))" fun addition_stage1 :: "nat \ nat list \ nat \ nat \ nat" where "addition_stage1 (as, lm) m p = (if as = 0 \ as = 1 \ as = 2 \ as = 3 then 2 else if as = 4 \ as = 5 \ as = 6 then 1 else 0)" fun addition_stage2 :: "nat \ nat list \ nat \ nat \ nat" where "addition_stage2 (as, lm) m p = (if 0 \ as \ as \ 3 then lm ! m else if 4 \ as \ as \ 6 then lm ! p else 0)" fun addition_stage3 :: "nat \ nat list \ nat \ nat \ nat" where "addition_stage3 (as, lm) m p = (if as = 1 then 4 else if as = 2 then 3 else if as = 3 then 2 else if as = 0 then 1 else if as = 5 then 2 else if as = 6 then 1 else if as = 4 then 0 else 0)" fun addition_measure :: "((nat \ nat list) \ nat \ nat) \ (nat \ nat \ nat)" where "addition_measure ((as, lm), m, p) = (addition_stage1 (as, lm) m p, addition_stage2 (as, lm) m p, addition_stage3 (as, lm) m p)" definition addition_LE :: "(((nat \ nat list) \ nat \ nat) \ ((nat \ nat list) \ nat \ nat)) set" where "addition_LE \ (inv_image lex_triple addition_measure)" lemma wf_additon_LE[simp]: "wf addition_LE" by(auto simp: addition_LE_def lex_triple_def lex_pair_def) declare addition_inv.simps[simp del] lemma update_zero_to_zero[simp]: "\am ! n = (0::nat); n < length am\ \ am[n := 0] = am" apply(simp add: list_update_same_conv) done lemma addition_inv_init: "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ addition_inv (0, lm) m n p lm" apply(simp add: addition_inv.simps Let_def ) apply(rule_tac x = "lm ! m" in exI, simp) done lemma abs_fetch[simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)" "abc_fetch (Suc 0) (addition m n p) = Some (Inc n)" "abc_fetch 2 (addition m n p) = Some (Inc p)" "abc_fetch 3 (addition m n p) = Some (Goto 0)" "abc_fetch 4 (addition m n p) = Some (Dec p 7)" "abc_fetch 5 (addition m n p) = Some (Inc m)" "abc_fetch 6 (addition m n p) = Some (Goto 4)" by(simp_all add: abc_fetch.simps addition.simps) lemma exists_small_list_elem1[simp]: "\m \ n; p < length lm; lm ! p = 0; m < p; n < p; x \ lm ! m; 0 < x\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xam \ n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\ \ \xa\lm ! m. lm[m := x, n := lm ! n + lm ! m - x, p := lm ! m - x] = lm[m := xa, n := lm ! n + lm ! m - xa, p := lm ! m - xa]" apply(rule_tac x = x in exI, simp) done lemma exists_small_list_elem5[simp]: "\m \ n; p < length lm; lm ! p = 0; m < p; n < p; x \ lm ! m; lm ! m \ x\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xa\lm ! m. lm[m := Suc x, n := lm ! n + lm ! m, p := lm ! m - Suc x] = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]" apply(rule_tac x = "Suc x" in exI, simp) done lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm" apply(cases asm, simp add: abc_steps_l.simps) done lemma list_double_update_2: "lm[a := x, b := y, a := z] = lm[b := y, a:=z]" by (metis list_update_overwrite list_update_swap) declare Let_def[simp] lemma addition_halt_lemma: "\m \ n; max m n < p; length lm > p\ \ \na. \ (\(as, lm') (m, p). as = 7) (abc_steps_l (0, lm) (addition m n p) na) (m, p) \ addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm \ addition_inv (abc_steps_l (0, lm) (addition m n p) (Suc na)) m n p lm \ ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), abc_steps_l (0, lm) (addition m n p) na, m, p) \ addition_LE" proof - assume assms_1: "m\n" and assms_2: "max m n < p" and assms_3: "length lm > p" { fix na obtain a b where ab:"abc_steps_l (0, lm) (addition m n p) na = (a, b)" by force assume assms2: "\ (\(as, lm') (m, p). as = 7) (abc_steps_l (0, lm) (addition m n p) na) (m, p)" "addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm" have r1:"addition_inv (abc_steps_l (0, lm) (addition m n p) (Suc na)) m n p lm" using assms_1 assms_2 assms_3 assms2 unfolding abc_step_red2 ab abc_step_l.simps abc_lm_v.simps abc_lm_s.simps addition_inv.simps by (auto split:if_splits simp add: addition_inv.simps Suc_diff_Suc) have r2:"((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), abc_steps_l (0, lm) (addition m n p) na, m, p) \ addition_LE" using assms_1 assms_2 assms_3 assms2 unfolding abc_step_red2 ab apply(auto split:if_splits simp add: addition_inv.simps abc_steps_zero) by(auto simp add: addition_LE_def lex_triple_def lex_pair_def abc_step_l.simps abc_lm_v.simps abc_lm_s.simps split: if_splits) note r1 r2 } thus ?thesis by auto qed lemma addition_correct': "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ \ stp. (\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) (abc_steps_l (0, lm) (addition m n p) stp)" apply(insert halt_lemma2[of addition_LE "\ ((as, lm'), m, p). addition_inv (as, lm') m n p lm" "\ stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)" "\ ((as, lm'), m, p). as = 7"], simp add: abc_steps_zero addition_inv_init) apply(drule_tac addition_halt_lemma,force,force) apply (simp,erule_tac exE) apply(rename_tac na) apply(rule_tac x = na in exI) apply(auto) done lemma length_addition[simp]: "length (addition a b c) = 7" by(auto simp: addition.simps) lemma addition_correct: assumes "m \ n" "max m n < p" "length lm > p" "lm ! p = 0" shows "\\ a. a = lm\ (addition m n p) \\ nl. addition_inv (7, nl) m n p lm\" using assms proof(rule_tac abc_Hoare_haltI, simp) fix lma assume "m \ n" "m < p \ n < p" "p < length lm" "lm ! p = 0" then have "\ stp. (\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) (abc_steps_l (0, lm) (addition m n p) stp)" by(rule_tac addition_correct', auto simp: addition_inv.simps) then obtain stp where "(\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) (abc_steps_l (0, lm) (addition m n p) stp)" using exE by presburger thus "\na. abc_final (abc_steps_l (0, lm) (addition m n p) na) (addition m n p) \ (\nl. addition_inv (7, nl) m n p lm) abc_holds_for abc_steps_l (0, lm) (addition m n p) na" by(auto intro:exI[of _ stp]) qed subsubsection \Correctness of compilation for constructor s\ lemma compile_s_correct': "\\nl. nl = n # 0 \ 2 @ anything\ addition 0 (Suc 0) 2 [+] [Inc (Suc 0)] \\nl. nl = n # Suc n # 0 # anything\" proof(rule_tac abc_Hoare_plus_halt) show "\\nl. nl = n # 0 \ 2 @ anything\ addition 0 (Suc 0) 2 \\ nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \ 2 @ anything)\" by(rule_tac addition_correct, auto simp: numeral_2_eq_2) next show "\\nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \ 2 @ anything)\ [Inc (Suc 0)] \\nl. nl = n # Suc n # 0 # anything\" by(rule_tac abc_Hoare_haltI, rule_tac x = 1 in exI, auto simp: addition_inv.simps abc_steps_l.simps abc_step_l.simps abc_fetch.simps numeral_2_eq_2 abc_lm_s.simps abc_lm_v.simps) qed declare rec_exec.simps[simp del] lemma abc_comp_commute: "(A [+] B) [+] C = A [+] (B [+] C)" apply(auto simp: abc_comp.simps abc_shift.simps) apply(rename_tac x) apply(case_tac x, auto) done lemma compile_s_correct: "\rec_ci s = (ap, arity, fp); rec_exec s [n] = r\ \ \\nl. nl = n # 0 \ (fp - arity) @ anything\ ap \\nl. nl = n # r # 0 \ (fp - Suc arity) @ anything\" apply(auto simp: rec_ci.simps rec_ci_s_def compile_s_correct' rec_exec.simps) done subsubsection \Correctness of compilation for constructor z\ lemma compile_z_correct: "\rec_ci z = (ap, arity, fp); rec_exec z [n] = r\ \ \\nl. nl = n # 0 \ (fp - arity) @ anything\ ap \\nl. nl = n # r # 0 \ (fp - Suc arity) @ anything\" apply(rule_tac abc_Hoare_haltI) apply(rule_tac x = 1 in exI) apply(auto simp: abc_steps_l.simps rec_ci.simps rec_ci_z_def numeral_2_eq_2 abc_fetch.simps abc_step_l.simps rec_exec.simps) done subsubsection \Correctness of compilation for constructor id\ lemma compile_id_correct': assumes "n < length args" shows "\\nl. nl = args @ 0 \ 2 @ anything\ addition n (length args) (Suc (length args)) \\nl. nl = args @ rec_exec (recf.id (length args) n) args # 0 # anything\" proof - have "\\nl. nl = args @ 0 \ 2 @ anything\ addition n (length args) (Suc (length args)) \\nl. addition_inv (7, nl) n (length args) (Suc (length args)) (args @ 0 \ 2 @ anything)\" using assms by(rule_tac addition_correct, auto simp: numeral_2_eq_2 nth_append) thus "?thesis" using assms by(simp add: addition_inv.simps rec_exec.simps nth_append numeral_2_eq_2 list_update_append) qed lemma compile_id_correct: "\n < m; length xs = m; rec_ci (recf.id m n) = (ap, arity, fp); rec_exec (recf.id m n) xs = r\ \ \\nl. nl = xs @ 0 \ (fp - arity) @ anything\ ap \\nl. nl = xs @ r # 0 \ (fp - Suc arity) @ anything\" apply(auto simp: rec_ci.simps rec_ci_id.simps compile_id_correct') done subsubsection \Correctness of compilation for constructor Cn\ lemma cn_merge_gs_tl_app: "cn_merge_gs (gs @ [g]) pstr = cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)" apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, auto) apply(simp add: abc_comp_commute) done lemma footprint_ge: "rec_ci a = (p, arity, fp) \ arity < fp" proof(induct a) case (Cn x1 a x3) then show ?case by(cases "rec_ci a", auto simp:rec_ci.simps) next case (Pr x1 a1 a2) then show ?case by(cases "rec_ci a1";cases "rec_ci a2", auto simp:rec_ci.simps) next case (Mn x1 a) then show ?case by(cases "rec_ci a", auto simp:rec_ci.simps) qed (auto simp: rec_ci.simps) lemma param_pattern: "\terminate f xs; rec_ci f = (p, arity, fp)\ \ length xs = arity" proof(induct arbitrary: p arity fp rule: terminate.induct) case (termi_cn f xs gs n) thus ?case by(cases "rec_ci f", (auto simp: rec_ci.simps)) next case (termi_pr x g xs n f) thus ?case by (cases "rec_ci f", cases "rec_ci g", auto simp: rec_ci.simps) next case (termi_mn xs n f r) thus ?case by (cases "rec_ci f", auto simp: rec_ci.simps) qed (auto simp: rec_ci.simps) lemma replicate_merge_anywhere: "x\a @ x\b @ ys = x\(a+b) @ ys" by(simp add:replicate_add) fun mv_box_inv :: "nat \ nat list \ nat \ nat \ nat list \ bool" where "mv_box_inv (as, lm) m n initlm = (let plus = initlm ! m + initlm ! n in length initlm > max m n \ m \ n \ (if as = 0 then \ k l. lm = initlm[m := k, n := l] \ k + l = plus \ k \ initlm ! m else if as = 1 then \ k l. lm = initlm[m := k, n := l] \ k + l + 1 = plus \ k < initlm ! m else if as = 2 then \ k l. lm = initlm[m := k, n := l] \ k + l = plus \ k \ initlm ! m else if as = 3 then lm = initlm[m := 0, n := plus] else False))" fun mv_box_stage1 :: "nat \ nat list \ nat \ nat" where "mv_box_stage1 (as, lm) m = (if as = 3 then 0 else 1)" fun mv_box_stage2 :: "nat \ nat list \ nat \ nat" where "mv_box_stage2 (as, lm) m = (lm ! m)" fun mv_box_stage3 :: "nat \ nat list \ nat \ nat" where "mv_box_stage3 (as, lm) m = (if as = 1 then 3 else if as = 2 then 2 else if as = 0 then 1 else 0)" fun mv_box_measure :: "((nat \ nat list) \ nat) \ (nat \ nat \ nat)" where "mv_box_measure ((as, lm), m) = (mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m, mv_box_stage3 (as, lm) m)" definition lex_pair :: "((nat \ nat) \ nat \ nat) set" where "lex_pair = less_than <*lex*> less_than" definition lex_triple :: "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" where "lex_triple \ less_than <*lex*> lex_pair" definition mv_box_LE :: "(((nat \ nat list) \ nat) \ ((nat \ nat list) \ nat)) set" where "mv_box_LE \ (inv_image lex_triple mv_box_measure)" lemma wf_lex_triple: "wf lex_triple" by (auto simp:lex_triple_def lex_pair_def) lemma wf_mv_box_le[intro]: "wf mv_box_LE" by(auto intro:wf_lex_triple simp: mv_box_LE_def) declare mv_box_inv.simps[simp del] lemma mv_box_inv_init: "\m < length initlm; n < length initlm; m \ n\ \ mv_box_inv (0, initlm) m n initlm" apply(simp add: abc_steps_l.simps mv_box_inv.simps) apply(rule_tac x = "initlm ! m" in exI, rule_tac x = "initlm ! n" in exI, simp) done lemma abc_fetch[simp]: "abc_fetch 0 (mv_box m n) = Some (Dec m 3)" "abc_fetch (Suc 0) (mv_box m n) = Some (Inc n)" "abc_fetch 2 (mv_box m n) = Some (Goto 0)" "abc_fetch 3 (mv_box m n) = None" apply(simp_all add: mv_box.simps abc_fetch.simps) done lemma replicate_Suc_iff_anywhere: "x # x\b @ ys = x\(Suc b) @ ys" by simp lemma exists_smaller_in_list0[simp]: "\m \ n; m < length initlm; n < length initlm; k + l = initlm ! m + initlm ! n; k \ initlm ! m; 0 < k\ \ \ka la. initlm[m := k, n := l, m := k - Suc 0] = initlm[m := ka, n := la] \ Suc (ka + la) = initlm ! m + initlm ! n \ ka < initlm ! m" apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, auto) apply(subgoal_tac "initlm[m := k, n := l, m := k - Suc 0] = initlm[n := l, m := k, m := k - Suc 0]",force intro:list_update_swap) by(simp add: list_update_swap) lemma exists_smaller_in_list1[simp]: "\m \ n; m < length initlm; n < length initlm; Suc (k + l) = initlm ! m + initlm ! n; k < initlm ! m\ \ \ka la. initlm[m := k, n := l, n := Suc l] = initlm[m := ka, n := la] \ ka + la = initlm ! m + initlm ! n \ ka \ initlm ! m" apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto) done lemma abc_steps_prop[simp]: "\length initlm > max m n; m \ n\ \ \ (\(as, lm) m. as = 3) (abc_steps_l (0, initlm) (mv_box m n) na) m \ mv_box_inv (abc_steps_l (0, initlm) (mv_box m n) na) m n initlm \ mv_box_inv (abc_steps_l (0, initlm) (mv_box m n) (Suc na)) m n initlm \ ((abc_steps_l (0, initlm) (mv_box m n) (Suc na), m), abc_steps_l (0, initlm) (mv_box m n) na, m) \ mv_box_LE" apply(rule impI, simp add: abc_step_red2) apply(cases "(abc_steps_l (0, initlm) (mv_box m n) na)", simp) apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps) apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def abc_step_l.simps abc_steps_l.simps mv_box_inv.simps abc_lm_v.simps abc_lm_s.simps split: if_splits ) apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp) done lemma mv_box_inv_halt: "\length initlm > max m n; m \ n\ \ \ stp. (\ (as, lm). as = 3 \ mv_box_inv (as, lm) m n initlm) (abc_steps_l (0::nat, initlm) (mv_box m n) stp)" apply(insert halt_lemma2[of mv_box_LE "\ ((as, lm), m). mv_box_inv (as, lm) m n initlm" "\ stp. (abc_steps_l (0, initlm) (mv_box m n) stp, m)" "\ ((as, lm), m). as = (3::nat)" ]) apply(insert wf_mv_box_le) apply(simp add: mv_box_inv_init abc_steps_zero) apply(erule_tac exE) by (metis (no_types, lifting) case_prodE' case_prodI) lemma mv_box_halt_cond: "\m \ n; mv_box_inv (a, b) m n lm; a = 3\ \ b = lm[n := lm ! m + lm ! n, m := 0]" apply(simp add: mv_box_inv.simps, auto) apply(simp add: list_update_swap) done lemma mv_box_correct': "\length lm > max m n; m \ n\ \ \ stp. abc_steps_l (0::nat, lm) (mv_box m n) stp = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])" by(drule mv_box_inv_halt, auto dest:mv_box_halt_cond) lemma length_mvbox[simp]: "length (mv_box m n) = 3" by(simp add: mv_box.simps) lemma mv_box_correct: "\length lm > max m n; m\n\ \ \\ nl. nl = lm\ mv_box m n \\ nl. nl = lm[n := (lm ! m + lm ! n), m:=0]\" apply(drule_tac mv_box_correct', simp) apply(auto simp: abc_Hoare_halt_def) by (metis abc_final.simps abc_holds_for.simps length_mvbox) declare list_update.simps(2)[simp del] lemma zero_case_rec_exec[simp]: "\length xs < gf; gf \ ft; n < length gs\ \ (rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) [ft + n - length xs := rec_exec (gs ! n) xs, 0 := 0] = 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything" using list_update_append[of "rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs)" "0 \ (length gs - n) @ 0 # 0 \ length xs @ anything" "ft + n - length xs" "rec_exec (gs ! n) xs"] apply(auto) apply(cases "length gs - n", simp, simp add: list_update.simps replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc) apply(simp add: list_update.simps) done lemma compile_cn_gs_correct': assumes g_cond: "\g\set (take n gs). terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. \\nl. nl = xs @ 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc\))" and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" shows "\\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci (take n gs)) ft \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0\(length gs - n) @ 0 \ Suc (length xs) @ anything\" using g_cond proof(induct n) case 0 have "ft > length xs" using ft by simp thus "?case" apply(rule_tac abc_Hoare_haltI) apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps replicate_add[THEN sym] replicate_Suc[THEN sym] del: replicate_Suc) done next case (Suc n) have ind': "\g\set (take n gs). terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. \\nl. nl = xs @ 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc\)) \ \\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci (take n gs)) ft \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything\" by fact have g_newcond: "\g\set (take (Suc n) gs). terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. \\nl. nl = xs @ 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc\))" by fact from g_newcond have ind: "\\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci (take n gs)) ft \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything\" apply(rule_tac ind', rule_tac ballI, erule_tac x = g in ballE, simp_all add: take_Suc) by(cases "n < length gs", simp add:take_Suc_conv_app_nth, simp) show "?case" proof(cases "n < length gs") case True have h: "n < length gs" by fact thus "?thesis" proof(simp add: take_Suc_conv_app_nth cn_merge_gs_tl_app) obtain gp ga gf where a: "rec_ci (gs!n) = (gp, ga, gf)" by (metis prod_cases3) moreover have "min (length gs) n = n" using h by simp moreover have "\\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci (take n gs)) ft [+] (gp [+] mv_box ga (ft + n)) \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything\" proof(rule_tac abc_Hoare_plus_halt) show "\\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci (take n gs)) ft \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything\" using ind by simp next have x: "gs!n \ set (take (Suc n) gs)" using h by(simp add: take_Suc_conv_app_nth) have b: "terminate (gs!n) xs" using a g_newcond h x by(erule_tac x = "gs!n" in ballE, simp_all) hence c: "length xs = ga" using a param_pattern by metis have d: "gf > ga" using footprint_ge a by simp have e: "ft \ gf" using ft a h Max_ge image_eqI by(simp, rule_tac max.coboundedI2, rule_tac Max_ge, simp, rule_tac insertI2, rule_tac f = "(\(aprog, p, n). n)" and x = "rec_ci (gs!n)" in image_eqI, simp, rule_tac x = "gs!n" in image_eqI, simp, simp) show "\\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything\ gp [+] mv_box ga (ft + n) \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything\" proof(rule_tac abc_Hoare_plus_halt) show "\\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything\ gp \\nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything\" proof - have "(\\nl. nl = xs @ 0 \ (gf - ga) @ 0\(ft - gf)@map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything\ gp \\nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \ (gf - Suc ga) @ 0\(ft - gf)@map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything\)" using a g_newcond h x apply(erule_tac x = "gs!n" in ballE) apply(simp, simp) done thus "?thesis" using a b c d e by(simp add: replicate_merge_anywhere) qed next show "\\nl. nl = xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything\ mv_box ga (ft + n) \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything\" proof - have "\\nl. nl = xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything\ mv_box ga (ft + n) \\nl. nl = (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! ga + (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! (ft + n), ga := 0]\" using a c d e h apply(rule_tac mv_box_correct) apply(simp_all) done moreover have "(xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! ga + (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! (ft + n), ga := 0]= xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything" using a c d e h by(simp add: list_update_append nth_append split: if_splits, auto) ultimately show "?thesis" by(simp) qed qed qed ultimately show "\\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci (take n gs)) ft [+] (case rec_ci (gs ! n) of (gprog, gpara, gn) \ gprog [+] mv_box gpara (ft + n)) \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything\" by simp qed next case False have h: "\ n < length gs" by fact hence ind': "\\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci gs) ft \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything\" using ind by simp thus "?thesis" using h by(simp) qed qed lemma compile_cn_gs_correct: assumes g_cond: "\g\set gs. terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. \\nl. nl = xs @ 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc\))" and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" shows "\\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything\ cn_merge_gs (map rec_ci gs) ft \\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything\" using assms using compile_cn_gs_correct'[of "length gs" gs xs ft ffp anything ] apply(auto) done lemma length_mvboxes[simp]: "length (mv_boxes aa ba n) = 3*n" by(induct n, auto simp: mv_boxes.simps) lemma exp_suc: "a\Suc b = a\b @ [a]" by(simp add: exp_ind del: replicate.simps) lemma last_0[simp]: "\Suc n \ ba - aa; length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\ \ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)" proof - assume h: "Suc n \ ba - aa" and g: "length lm2 = Suc n" "length lm3 = ba - Suc (aa + n)" from h and g have k: "ba - aa = Suc (length lm3 + n)" by arith from k show "(last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = 0" apply(simp, insert g) apply(simp add: nth_append) done qed lemma butlast_last[simp]: "length lm1 = aa \ (lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2" apply(simp add: nth_append) done lemma arith_as_simp[simp]: "\Suc n \ ba - aa; aa < ba\ \ (ba < Suc (aa + (ba - Suc (aa + n) + n))) = False" apply arith done lemma butlast_elem[simp]: "\Suc n \ ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\ \ (lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0" using nth_append[of "lm1 @ (0::'a)\n @ last lm2 # lm3 @ butlast lm2" "(0::'a) # lm4" "ba + n"] apply(simp) done lemma update_butlast_eq0[simp]: "\Suc n \ ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\ \ (lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4) [ba + n := last lm2, aa + n := 0] = lm1 @ 0 # 0\n @ lm3 @ lm2 @ lm4" using list_update_append[of "lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2" "0 # lm4" "ba + n" "last lm2"] apply(simp add: list_update_append list_update.simps(2-) replicate_Suc_iff_anywhere exp_suc del: replicate_Suc) apply(cases lm2, simp, simp) done lemma update_butlast_eq1[simp]: "\Suc (length lm1 + n) \ ba; length lm2 = Suc n; length lm3 = ba - Suc (length lm1 + n); \ ba - Suc (length lm1) < ba - Suc (length lm1 + n); \ ba + n - length lm1 < n\ \ (0::nat) \ n @ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4)[ba - length lm1 := last lm2, 0 := 0] = 0 # 0 \ n @ lm3 @ lm2 @ lm4" apply(subgoal_tac "ba - length lm1 = Suc n + length lm3", simp add: list_update.simps(2-) list_update_append) apply(simp add: replicate_Suc_iff_anywhere exp_suc del: replicate_Suc) apply(cases lm2, simp, simp) apply(auto) done lemma mv_boxes_correct: "\aa + n \ ba; ba > aa; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\ \ \\ nl. nl = lm1 @ lm2 @ lm3 @ 0\n @ lm4\ (mv_boxes aa ba n) \\ nl. nl = lm1 @ 0\n @ lm3 @ lm2 @ lm4\" proof(induct n arbitrary: lm2 lm3 lm4) case 0 thus "?case" by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) next case (Suc n) have ind: "\lm2 lm3 lm4. \aa + n \ ba; aa < ba; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\ \ \\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ n @ lm4\ mv_boxes aa ba n \\nl. nl = lm1 @ 0 \ n @ lm3 @ lm2 @ lm4\" by fact have h1: "aa + Suc n \ ba" by fact have h2: "aa < ba" by fact have h3: "length lm1 = aa" by fact have h4: "length lm2 = Suc n" by fact have h5: "length lm3 = ba - aa - Suc n" by fact have "\\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4\ mv_boxes aa ba n [+] mv_box (aa + n) (ba + n) \\nl. nl = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4\" proof(rule_tac abc_Hoare_plus_halt) have "\\nl. nl = lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \ n @ (0 # lm4)\ mv_boxes aa ba n \\ nl. nl = lm1 @ 0\n @ (last lm2 # lm3) @ butlast lm2 @ (0 # lm4)\" using h1 h2 h3 h4 h5 by(rule_tac ind, simp_all) moreover have " lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \ n @ (0 # lm4) = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4" using h4 by(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc, cases lm2, simp_all) ultimately show "\\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4\ mv_boxes aa ba n \\ nl. nl = lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4\" by (metis append_Cons) next let ?lm = "lm1 @ 0 \ n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4" have "\\nl. nl = ?lm\ mv_box (aa + n) (ba + n) \\ nl. nl = ?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0]\" using h1 h2 h3 h4 h5 by(rule_tac mv_box_correct, simp_all) moreover have "?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0] = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4" using h1 h2 h3 h4 h5 by(auto simp: nth_append list_update_append split: if_splits) ultimately show "\\nl. nl = lm1 @ 0 \ n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4\ mv_box (aa + n) (ba + n) \\nl. nl = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4\" by simp qed thus "?case" by(simp add: mv_boxes.simps) qed lemma update_butlast_eq2[simp]: "\Suc n \ aa - length lm1; length lm1 < aa; length lm2 = aa - Suc (length lm1 + n); length lm3 = Suc n; \ aa - Suc (length lm1) < aa - Suc (length lm1 + n); \ aa + n - length lm1 < n\ \ butlast lm3 @ ((0::nat) # lm2 @ 0 \ n @ last lm3 # lm4)[0 := last lm3, aa - length lm1 := 0] = lm3 @ lm2 @ 0 # 0 \ n @ lm4" apply(subgoal_tac "aa - length lm1 = length lm2 + Suc n") apply(simp add: list_update.simps list_update_append) apply(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc) apply(cases lm3, simp, simp) apply(auto) done lemma mv_boxes_correct2: "\n \ aa - ba; ba < aa; length (lm1::nat list) = ba; length (lm2::nat list) = aa - ba - n; length (lm3::nat list) = n\ \\\ nl. nl = lm1 @ 0\n @ lm2 @ lm3 @ lm4\ (mv_boxes aa ba n) \\ nl. nl = lm1 @ lm3 @ lm2 @ 0\n @ lm4\" proof(induct n arbitrary: lm2 lm3 lm4) case 0 thus "?case" by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) next case (Suc n) have ind: "\lm2 lm3 lm4. \n \ aa - ba; ba < aa; length lm1 = ba; length lm2 = aa - ba - n; length lm3 = n\ \ \\nl. nl = lm1 @ 0 \ n @ lm2 @ lm3 @ lm4\ mv_boxes aa ba n \\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ n @ lm4\" by fact have h1: "Suc n \ aa - ba" by fact have h2: "ba < aa" by fact have h3: "length lm1 = ba" by fact have h4: "length lm2 = aa - ba - Suc n" by fact have h5: "length lm3 = Suc n" by fact have "\\nl. nl = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4\ mv_boxes aa ba n [+] mv_box (aa + n) (ba + n) \\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4\" proof(rule_tac abc_Hoare_plus_halt) have "\\ nl. nl = lm1 @ 0 \ n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4)\ mv_boxes aa ba n \\ nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\n @ (last lm3 # lm4)\" using h1 h2 h3 h4 h5 by(rule_tac ind, simp_all) moreover have "lm1 @ 0 \ n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4) = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4" using h5 by(simp add: replicate_Suc_iff_anywhere exp_suc del: replicate_Suc, cases lm3, simp_all) ultimately show "\\nl. nl = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4\ mv_boxes aa ba n \\ nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\n @ (last lm3 # lm4)\" by metis next thm mv_box_correct let ?lm = "lm1 @ butlast lm3 @ (0 # lm2) @ 0 \ n @ last lm3 # lm4" have "\\nl. nl = ?lm\ mv_box (aa + n) (ba + n) \\nl. nl = ?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0]\" using h1 h2 h3 h4 h5 by(rule_tac mv_box_correct, simp_all) moreover have "?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0] = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4" using h1 h2 h3 h4 h5 by(auto simp: nth_append list_update_append split: if_splits) ultimately show "\\nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0 \ n @ last lm3 # lm4\ mv_box (aa + n) (ba + n) \\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4\" by simp qed thus "?case" by(simp add: mv_boxes.simps) qed lemma save_paras: "\\nl. nl = xs @ 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything\ mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) \\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything\" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" have "\\nl. nl = [] @ xs @ (0\(?ft - length xs) @ map (\i. rec_exec i xs) gs @ [0]) @ 0 \ (length xs) @ anything\ mv_boxes 0 (Suc ?ft + length gs) (length xs) \\nl. nl = [] @ 0 \ (length xs) @ (0\(?ft - length xs) @ map (\i. rec_exec i xs) gs @ [0]) @ xs @ anything\" by(rule_tac mv_boxes_correct, auto) thus "?thesis" by(simp add: replicate_merge_anywhere) qed lemma length_le_max_insert_rec_ci[intro]: "length gs \ ffp \ length gs \ max x1 (Max (insert ffp (x2 ` x3 ` set gs)))" apply(rule_tac max.coboundedI2) apply(simp add: Max_ge_iff) done lemma restore_new_paras: "ffp \ length gs \ \\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything\ mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) \\nl. nl = map (\i. rec_exec i xs) gs @ 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ 0 # xs @ anything\" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" assume j: "ffp \ length gs" hence "\\ nl. nl = [] @ 0\length gs @ 0\(?ft - length gs) @ map (\i. rec_exec i xs) gs @ ((0 # xs) @ anything)\ mv_boxes ?ft 0 (length gs) \\ nl. nl = [] @ map (\i. rec_exec i xs) gs @ 0\(?ft - length gs) @ 0\length gs @ ((0 # xs) @ anything)\" by(rule_tac mv_boxes_correct2, auto) moreover have "?ft \ length gs" using j by(auto) ultimately show "?thesis" using j by(simp add: replicate_merge_anywhere le_add_diff_inverse) qed lemma le_max_insert[intro]: "ffp \ max x0 (Max (insert ffp (x1 ` x2 ` set gs)))" by (rule max.coboundedI2) auto declare max_less_iff_conj[simp del] lemma save_rs: "\far = length gs; ffp \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))); far < ffp\ \ \\nl. nl = map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ xs @ anything\ mv_box far (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) \\nl. nl = map (\i. rec_exec i xs) gs @ 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything\" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" thm mv_box_correct let ?lm= " map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ ?ft @ xs @ anything" assume h: "far = length gs" "ffp \ ?ft" "far < ffp" hence "\\ nl. nl = ?lm\ mv_box far ?ft \\ nl. nl = ?lm[?ft := ?lm!far + ?lm!?ft, far := 0]\" apply(rule_tac mv_box_correct) by( auto) moreover have "?lm[?ft := ?lm!far + ?lm!?ft, far := 0] = map (\i. rec_exec i xs) gs @ 0 \ (?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" using h apply(simp add: nth_append) using list_update_length[of "map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - Suc (length gs))" 0 "0 \ length gs @ xs @ anything" "rec_exec (Cn (length xs) f gs) xs"] apply(simp add: replicate_merge_anywhere replicate_Suc_iff_anywhere del: replicate_Suc) by(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere del: replicate_Suc) ultimately show "?thesis" by(simp) qed lemma length_empty_boxes[simp]: "length (empty_boxes n) = 2*n" apply(induct n, simp, simp) done lemma empty_one_box_correct: "\\nl. nl = 0 \ n @ x # lm\ [Dec n 2, Goto 0] \\nl. nl = 0 # 0 \ n @ lm\" proof(induct x) case 0 thus "?case" by(simp add: abc_Hoare_halt_def, rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps replicate_Suc[THEN sym] exp_suc del: replicate_Suc) next case (Suc x) have "\\nl. nl = 0 \ n @ x # lm\ [Dec n 2, Goto 0] \\nl. nl = 0 # 0 \ n @ lm\" by fact then obtain stp where "abc_steps_l (0, 0 \ n @ x # lm) [Dec n 2, Goto 0] stp = (Suc (Suc 0), 0 # 0 \ n @ lm)" apply(auto simp: abc_Hoare_halt_def) - by (smt abc_final.simps abc_holds_for.elims(2) length_Cons list.size(3)) + by (smt (verit) abc_final.simps abc_holds_for.elims(2) length_Cons list.size(3)) moreover have "abc_steps_l (0, 0\n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0)) = (0, 0 \ n @ x # lm)" by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps list_update.simps list_update_append) ultimately have "abc_steps_l (0, 0\n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0) + stp) = (Suc (Suc 0), 0 # 0\n @ lm)" by(simp only: abc_steps_add) thus "?case" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = "Suc (Suc stp)" in exI, simp) done qed lemma empty_boxes_correct: "length lm \ n \ \\ nl. nl = lm\ empty_boxes n \\ nl. nl = 0\n @ drop n lm\" proof(induct n) case 0 thus "?case" by(simp add: empty_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) next case (Suc n) have ind: "n \ length lm \ \\nl. nl = lm\ empty_boxes n \\nl. nl = 0 \ n @ drop n lm\" by fact have h: "Suc n \ length lm" by fact have "\\nl. nl = lm\ empty_boxes n [+] [Dec n 2, Goto 0] \\nl. nl = 0 # 0 \ n @ drop (Suc n) lm\" proof(rule_tac abc_Hoare_plus_halt) show "\\nl. nl = lm\ empty_boxes n \\nl. nl = 0 \ n @ drop n lm\" using h by(rule_tac ind, simp) next show "\\nl. nl = 0 \ n @ drop n lm\ [Dec n 2, Goto 0] \\nl. nl = 0 # 0 \ n @ drop (Suc n) lm\" using empty_one_box_correct[of n "lm ! n" "drop (Suc n) lm"] using h by(simp add: Cons_nth_drop_Suc) qed thus "?case" by(simp add: empty_boxes.simps) qed lemma insert_dominated[simp]: "length gs \ ffp \ length gs + (max xs (Max (insert ffp (x1 ` x2 ` set gs))) - length gs) = max xs (Max (insert ffp (x1 ` x2 ` set gs)))" apply(rule_tac le_add_diff_inverse) apply(rule_tac max.coboundedI2) apply(simp add: Max_ge_iff) done lemma clean_paras: "ffp \ length gs \ \\nl. nl = map (\i. rec_exec i xs) gs @ 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything\ empty_boxes (length gs) \\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything\" proof- let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" assume h: "length gs \ ffp" let ?lm = "map (\i. rec_exec i xs) gs @ 0 \ (?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" have "\\ nl. nl = ?lm\ empty_boxes (length gs) \\ nl. nl = 0\length gs @ drop (length gs) ?lm\" by(rule_tac empty_boxes_correct, simp) moreover have "0\length gs @ drop (length gs) ?lm = 0 \ ?ft @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" using h by(simp add: replicate_merge_anywhere) ultimately show "?thesis" by metis qed lemma restore_rs: "\\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything\ mv_box (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) \\nl. nl = 0 \ length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - (length xs)) @ 0 \ length gs @ xs @ anything\" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" let ?lm = "0\(length xs) @ 0\(?ft - (length xs)) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" thm mv_box_correct have "\\ nl. nl = ?lm\ mv_box ?ft (length xs) \\ nl. nl = ?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]\" by(rule_tac mv_box_correct, simp, simp) moreover have "?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0] = 0 \ length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - (length xs)) @ 0 \ length gs @ xs @ anything" apply(auto simp: list_update_append nth_append) (* ~ 7 sec *) apply(cases ?ft, simp_all add: Suc_diff_le list_update.simps) apply(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc) done ultimately show "?thesis" by(simp add: replicate_merge_anywhere) qed lemma restore_orgin_paras: "\\nl. nl = 0 \ length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ 0 \ length gs @ xs @ anything\ mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs) \\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything\" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" thm mv_boxes_correct2 have "\\ nl. nl = [] @ 0\(length xs) @ (rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - length xs) @ 0 \ length gs) @ xs @ anything\ mv_boxes (Suc ?ft + length gs) 0 (length xs) \\ nl. nl = [] @ xs @ (rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - length xs) @ 0 \ length gs) @ 0\length xs @ anything\" by(rule_tac mv_boxes_correct2, auto) thus "?thesis" by(simp add: replicate_merge_anywhere) qed lemma compile_cn_correct': assumes f_ind: "\ anything r. rec_exec f (map (\g. rec_exec g xs) gs) = rec_exec (Cn (length xs) f gs) xs \ \\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ anything\ fap \\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (ffp - Suc far) @ anything\" and compile: "rec_ci f = (fap, far, ffp)" and term_f: "terminate f (map (\g. rec_exec g xs) gs)" and g_cond: "\g\set gs. terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. \\nl. nl = xs @ 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc\))" shows "\\nl. nl = xs @ 0 # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything\ cn_merge_gs (map rec_ci gs) (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) [+] (mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) [+] (mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) [+] (fap [+] (mv_box far (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) [+] (empty_boxes (length gs) [+] (mv_box (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) [+] mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs))))))) \\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything\" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" let ?A = "cn_merge_gs (map rec_ci gs) ?ft" let ?B = "mv_boxes 0 (Suc (?ft+length gs)) (length xs)" let ?C = "mv_boxes ?ft 0 (length gs)" let ?D = fap let ?E = "mv_box far ?ft" let ?F = "empty_boxes (length gs)" let ?G = "mv_box ?ft (length xs)" let ?H = "mv_boxes (Suc (?ft + length gs)) 0 (length xs)" let ?P1 = "\nl. nl = xs @ 0 # 0 \ (?ft + length gs) @ anything" let ?S = "\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft + length gs) @ anything" let ?Q1 = "\ nl. nl = xs @ 0\(?ft - length xs) @ map (\ i. rec_exec i xs) gs @ 0\(Suc (length xs)) @ anything" show "\?P1\ (?A [+] (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))))) \?S\" proof(rule_tac abc_Hoare_plus_halt) show "\?P1\ ?A \?Q1\" using g_cond by(rule_tac compile_cn_gs_correct, auto) next let ?Q2 = "\nl. nl = 0 \ ?ft @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything" show "\?Q1\ (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H)))))) \?S\" proof(rule_tac abc_Hoare_plus_halt) show "\?Q1\ ?B \?Q2\" by(rule_tac save_paras) next let ?Q3 = "\ nl. nl = map (\i. rec_exec i xs) gs @ 0\?ft @ 0 # xs @ anything" show "\?Q2\ (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))) \?S\" proof(rule_tac abc_Hoare_plus_halt) have "ffp \ length gs" using compile term_f apply(subgoal_tac "length gs = far") apply(drule_tac footprint_ge, simp) by(drule_tac param_pattern, auto) thus "\?Q2\ ?C \?Q3\" by(erule_tac restore_new_paras) next let ?Q4 = "\ nl. nl = map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0\?ft @ xs @ anything" have a: "far = length gs" using compile term_f by(drule_tac param_pattern, auto) have b:"?ft \ ffp" by auto have c: "ffp > far" using compile by(erule_tac footprint_ge) show "\?Q3\ (?D [+] (?E [+] (?F [+] (?G [+] ?H)))) \?S\" proof(rule_tac abc_Hoare_plus_halt) have "\\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ 0\(?ft - ffp + far) @ 0 # xs @ anything\ fap \\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (ffp - Suc far) @ 0\(?ft - ffp + far) @ 0 # xs @ anything\" by(rule_tac f_ind, simp add: rec_exec.simps) thus "\?Q3\ fap \?Q4\" using a b c by(simp add: replicate_merge_anywhere, cases "?ft", simp_all add: exp_suc del: replicate_Suc) next let ?Q5 = "\nl. nl = map (\i. rec_exec i xs) gs @ 0\(?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0\(length gs)@ xs @ anything" show "\?Q4\ (?E [+] (?F [+] (?G [+] ?H))) \?S\" proof(rule_tac abc_Hoare_plus_halt) from a b c show "\?Q4\ ?E \?Q5\" by(erule_tac save_rs, simp_all) next let ?Q6 = "\nl. nl = 0\?ft @ rec_exec (Cn (length xs) f gs) xs # 0\(length gs)@ xs @ anything" show "\?Q5\ (?F [+] (?G [+] ?H)) \?S\" proof(rule_tac abc_Hoare_plus_halt) have "length gs \ ffp" using a b c by simp thus "\?Q5\ ?F \?Q6\" by(erule_tac clean_paras) next let ?Q7 = "\nl. nl = 0\length xs @ rec_exec (Cn (length xs) f gs) xs # 0\(?ft - (length xs)) @ 0\(length gs)@ xs @ anything" show "\?Q6\ (?G [+] ?H) \?S\" proof(rule_tac abc_Hoare_plus_halt) show "\?Q6\ ?G \?Q7\" by(rule_tac restore_rs) next show "\?Q7\ ?H \?S\" by(rule_tac restore_orgin_paras) qed qed qed qed qed qed qed qed lemma compile_cn_correct: assumes termi_f: "terminate f (map (\g. rec_exec g xs) gs)" and f_ind: "\ap arity fp anything. rec_ci f = (ap, arity, fp) \ \\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (fp - arity) @ anything\ ap \\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec f (map (\g. rec_exec g xs) gs) # 0 \ (fp - Suc arity) @ anything\" and g_cond: "\g\set gs. terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. \\nl. nl = xs @ 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc\))" and compile: "rec_ci (Cn n f gs) = (ap, arity, fp)" and len: "length xs = n" shows "\\nl. nl = xs @ 0 \ (fp - arity) @ anything\ ap \\nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \ (fp - Suc arity) @ anything\" proof(cases "rec_ci f") fix fap far ffp assume h: "rec_ci f = (fap, far, ffp)" then have f_newind: "\ anything .\\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ anything\ fap \\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec f (map (\g. rec_exec g xs) gs) # 0 \ (ffp - Suc far) @ anything\" by(rule_tac f_ind, simp_all) thus "\\nl. nl = xs @ 0 \ (fp - arity) @ anything\ ap \\nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \ (fp - Suc arity) @ anything\" using compile len h termi_f g_cond apply(auto simp: rec_ci.simps abc_comp_commute) apply(rule_tac compile_cn_correct', simp_all) done qed subsubsection \Correctness of compilation for constructor Pr\ lemma mv_box_correct_simp[simp]: "\length xs = n; ft = max (n+3) (max fft gft)\ \ \\nl. nl = xs @ 0 # 0 \ (ft - n) @ anything\ mv_box n ft \\nl. nl = xs @ 0 # 0 \ (ft - n) @ anything\" using mv_box_correct[of n ft "xs @ 0 # 0 \ (ft - n) @ anything"] by(auto) lemma length_under_max[simp]: "length xs < max (length xs + 3) fft" by auto lemma save_init_rs: "\length xs = n; ft = max (n+3) (max fft gft)\ \ \\nl. nl = xs @ rec_exec f xs # 0 \ (ft - n) @ anything\ mv_box n (Suc n) \\nl. nl = xs @ 0 # rec_exec f xs # 0 \ (ft - Suc n) @ anything\" using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (ft - n) @ anything"] apply(auto simp: list_update_append list_update.simps nth_append split: if_splits) apply(cases "(max (length xs + 3) (max fft gft))", simp_all add: list_update.simps Suc_diff_le) done lemma less_then_max_plus2[simp]: "n + (2::nat) < max (n + 3) x" by auto lemma less_then_max_plus3[simp]: "n < max (n + (3::nat)) x" by auto lemma mv_box_max_plus_3_correct[simp]: "length xs = n \ \\nl. nl = xs @ x # 0 \ (max (n + (3::nat)) (max fft gft) - n) @ anything\ mv_box n (max (n + 3) (max fft gft)) \\nl. nl = xs @ 0 \ (max (n + 3) (max fft gft) - n) @ x # anything\" proof - assume h: "length xs = n" let ?ft = "max (n+3) (max fft gft)" let ?lm = "xs @ x # 0\(?ft - Suc n) @ 0 # anything" have g: "?ft > n + 2" by simp thm mv_box_correct have a: "\\ nl. nl = ?lm\ mv_box n ?ft \\ nl. nl = ?lm[?ft := ?lm!n + ?lm!?ft, n := 0]\" using h by(rule_tac mv_box_correct, auto) have b:"?lm = xs @ x # 0 \ (max (n + 3) (max fft gft) - n) @ anything" by(cases ?ft, simp_all add: Suc_diff_le exp_suc del: replicate_Suc) have c: "?lm[?ft := ?lm!n + ?lm!?ft, n := 0] = xs @ 0 \ (max (n + 3) (max fft gft) - n) @ x # anything" using h g apply(auto simp: nth_append list_update_append split: if_splits) using list_update_append[of "x # 0 \ (max (length xs + 3) (max fft gft) - Suc (length xs))" "0 # anything" "max (length xs + 3) (max fft gft) - length xs" "x"] apply(auto simp: if_splits) apply(simp add: list_update.simps replicate_Suc[THEN sym] del: replicate_Suc) done from a c show "?thesis" using h apply(simp) using b by simp qed lemma max_less_suc_suc[simp]: "max n (Suc n) < Suc (Suc (max (n + 3) x + anything - Suc 0))" by arith lemma suc_less_plus_3[simp]: "Suc n < max (n + 3) x" by arith lemma mv_box_ok_suc_simp[simp]: "length xs = n \ \\nl. nl = xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything\ mv_box n (Suc n) \\nl. nl = xs @ 0 # rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ x # anything\" using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything"] apply(simp add: nth_append list_update_append list_update.simps) apply(cases "max (n + 3) (max fft gft)", simp_all) apply(cases "max (n + 3) (max fft gft) - 1", simp_all add: Suc_diff_le list_update.simps(2)) done lemma abc_append_first_steps_eq_pre: assumes notfinal: "abc_notfinal (abc_steps_l (0, lm) A n) A" and notnull: "A \ []" shows "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" using notfinal proof(induct n) case 0 thus "?case" by(simp add: abc_steps_l.simps) next case (Suc n) have ind: "abc_notfinal (abc_steps_l (0, lm) A n) A \ abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" by fact have h: "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A" by fact then have a: "abc_notfinal (abc_steps_l (0, lm) A n) A" by(simp add: notfinal_Suc) then have b: "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" using ind by simp obtain s lm' where c: "abc_steps_l (0, lm) A n = (s, lm')" by (metis prod.exhaust) then have d: "s < length A \ abc_steps_l (0, lm) (A @ B) n = (s, lm')" using a b by simp thus "?case" using c by(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps nth_append) qed lemma abc_append_first_step_eq_pre: "st < length A \ abc_step_l (st, lm) (abc_fetch st (A @ B)) = abc_step_l (st, lm) (abc_fetch st A)" by(simp add: abc_step_l.simps abc_fetch.simps nth_append) lemma abc_append_first_steps_halt_eq': assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" and notnull: "A \ []" shows "\ n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')" proof - have "\ n'. abc_notfinal (abc_steps_l (0, lm) A n') A \ abc_final (abc_steps_l (0, lm) A (Suc n')) A" using assms by(rule_tac n = n in abc_before_final, simp_all) then obtain na where a: "abc_notfinal (abc_steps_l (0, lm) A na) A \ abc_final (abc_steps_l (0, lm) A (Suc na)) A" .. obtain sa lma where b: "abc_steps_l (0, lm) A na = (sa, lma)" by (metis prod.exhaust) then have c: "abc_steps_l (0, lm) (A @ B) na = (sa, lma)" using a abc_append_first_steps_eq_pre[of lm A na B] assms by simp have d: "sa < length A" using b a by simp then have e: "abc_step_l (sa, lma) (abc_fetch sa (A @ B)) = abc_step_l (sa, lma) (abc_fetch sa A)" by(rule_tac abc_append_first_step_eq_pre) from a have "abc_steps_l (0, lm) A (Suc na) = (length A, lm')" using final equal_when_halt by(cases "abc_steps_l (0, lm) A (Suc na)" , simp) then have "abc_steps_l (0, lm) (A @ B) (Suc na) = (length A, lm')" using a b c e by(simp add: abc_step_red2) thus "?thesis" by blast qed lemma abc_append_first_steps_halt_eq: assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" shows "\ n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')" using final apply(cases "A = []") apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps abc_exec_null) apply(rule_tac abc_append_first_steps_halt_eq', simp_all) done lemma suc_suc_max_simp[simp]: "Suc (Suc (max (xs + 3) fft - Suc (Suc ( xs)))) = max ( xs + 3) fft - ( xs)" by arith lemma contract_dec_ft_length_plus_7[simp]: "\ft = max (n + 3) (max fft gft); length xs = n\ \ \\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything\ [Dec ft (length gap + 7)] \\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything\" apply (simp add: abc_Hoare_halt_def) apply (rule_tac x = 1 in exI) apply (auto simp add: max_def) apply(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps list_update_append) using list_update_length [of "(x - Suc y) # rec_exec (Pr (length xs) f g) (xs @ [x - Suc y]) # 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" "Suc y" anything y] apply (auto simp add: Suc_diff_Suc) using numeral_3_eq_3 apply presburger using numeral_3_eq_3 apply presburger done lemma adjust_paras': "length xs = n \ \\nl. nl = xs @ x # y # anything\ [Inc n] [+] [Dec (Suc n) 2, Goto 0] \\nl. nl = xs @ Suc x # 0 # anything\" proof(rule_tac abc_Hoare_plus_halt) assume "length xs = n" thus "\\nl. nl = xs @ x # y # anything\ [Inc n] \\ nl. nl = xs @ Suc x # y # anything\" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = 1 in exI, force simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2)) done next assume h: "length xs = n" thus "\\nl. nl = xs @ Suc x # y # anything\ [Dec (Suc n) 2, Goto 0] \\nl. nl = xs @ Suc x # 0 # anything\" proof(induct y) case 0 thus "?case" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2)) done next case (Suc y) have "length xs = n \ \\nl. nl = xs @ Suc x # y # anything\ [Dec (Suc n) 2, Goto 0] \\nl. nl = xs @ Suc x # 0 # anything\" by fact then obtain stp where "abc_steps_l (0, xs @ Suc x # y # anything) [Dec (Suc n) 2, Goto 0] stp = (2, xs @ Suc x # 0 # anything)" using h apply(auto simp: abc_Hoare_halt_def numeral_2_eq_2) by (metis (mono_tags, lifting) abc_final.simps abc_holds_for.elims(2) length_Cons list.size(3)) moreover have "abc_steps_l (0, xs @ Suc x # Suc y # anything) [Dec (Suc n) 2, Goto 0] 2 = (0, xs @ Suc x # y # anything)" using h by(simp add: abc_steps_l.simps numeral_2_eq_2 abc_step_l.simps abc_fetch.simps abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2)) ultimately show "?case" apply(simp add: abc_Hoare_halt_def) by(rule exI[of _ "2 + stp"], simp only: abc_steps_add, simp) qed qed lemma adjust_paras: "length xs = n \ \\nl. nl = xs @ x # y # anything\ [Inc n, Dec (Suc n) 3, Goto (Suc 0)] \\nl. nl = xs @ Suc x # 0 # anything\" using adjust_paras'[of xs n x y anything] by(simp add: abc_comp.simps abc_shift.simps numeral_2_eq_2 numeral_3_eq_3) lemma rec_ci_SucSuc_n[simp]: "\rec_ci g = (gap, gar, gft); \yx\ \ gar = Suc (Suc n)" by(auto dest:param_pattern elim!:allE[of _ y]) lemma loop_back': assumes h: "length A = length gap + 4" "length xs = n" and le: "y \ x" shows "\ stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (length A, xs @ m # y # 0 # anything)" using le proof(induct x) case 0 thus "?case" using h by(rule_tac x = 0 in exI, auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_s.simps abc_lm_v.simps) next case (Suc x) have "x \ y \ \stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (length A, xs @ m # y # 0 # anything)" by fact moreover have "Suc x \ y" by fact moreover then have "\ stp. abc_steps_l (length A, xs @ m # (y - Suc x) # Suc x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (length A, xs @ m # (y - x) # x # anything)" using h apply(rule_tac x = 3 in exI) by(simp add: abc_steps_l.simps numeral_3_eq_3 abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps list_update_append list_update.simps(2)) ultimately show "?case" apply(auto simp add: abc_steps_add) by (metis abc_steps_add) qed lemma loop_back: assumes h: "length A = length gap + 4" "length xs = n" shows "\ stp. abc_steps_l (length A, xs @ m # 0 # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (0, xs @ m # x # 0 # anything)" using loop_back'[of A gap xs n x x m anything] assms apply(auto) apply(rename_tac stp) apply(rule_tac x = "stp + 1" in exI) apply(simp only: abc_steps_add, simp) apply(simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps) done lemma rec_exec_pr_0_simps: "rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs" by(simp add: rec_exec.simps) lemma rec_exec_pr_Suc_simps: "rec_exec (Pr n f g) (xs @ [Suc y]) = rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" apply(induct y) apply(simp add: rec_exec.simps) apply(simp add: rec_exec.simps) done lemma suc_max_simp[simp]: "Suc (max (n + 3) fft - Suc (Suc (Suc n))) = max (n + 3) fft - Suc (Suc n)" by arith lemma pr_loop: assumes code: "code = ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" and len: "length xs = n" and g_ind: "\ yanything. \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything\ gap \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything\)" and compile_g: "rec_ci g = (gap, gar, gft)" and termi_g: "\ y x" shows "\stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) code stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (ft - Suc (Suc n)) @ y # anything)" proof - let ?A = "[Dec ft (length gap + 7)]" let ?B = "gap" let ?C = "[Inc n, Dec (Suc n) 3, Goto (Suc 0)]" let ?D = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" have "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ((?A [+] (?B [+] ?C)) @ ?D) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" proof - have "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ((?A [+] (?B [+] ?C))) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" proof - have "\\ nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything\ (?A [+] (?B [+] ?C)) \\ nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything\" proof(rule_tac abc_Hoare_plus_halt) show "\\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything\ [Dec ft (length gap + 7)] \\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything\" using ft len by(simp) next show "\\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything\ ?B [+] ?C \\nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything\" proof(rule_tac abc_Hoare_plus_halt) have a: "gar = Suc (Suc n)" using compile_g termi_g len less by simp have b: "gft > gar" using compile_g by(erule_tac footprint_ge) show "\\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything\ gap \\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything\" proof - have "\\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (gft - gar) @ 0\(ft - gft) @ y # anything\ gap \\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # rec_exec g (xs @ [(x - Suc y), rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (gft - Suc gar) @ 0\(ft - gft) @ y # anything\" using g_ind less by simp thus "?thesis" using a b ft by(simp add: replicate_merge_anywhere numeral_3_eq_3) qed next show "\\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything\ [Inc n, Dec (Suc n) 3, Goto (Suc 0)] \\nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything\" using len less using adjust_paras[of xs n "x - Suc y" " rec_exec (Pr n f g) (xs @ [x - Suc y])" " rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything"] by(simp add: Suc_diff_Suc) qed qed thus "?thesis" apply(simp add: abc_Hoare_halt_def, auto) apply(rename_tac na) apply(rule_tac x = na in exI, case_tac "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ([Dec ft (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) na", simp) done qed then obtain stpa where "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ((?A [+] (?B [+] ?C))) stpa = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" .. thus "?thesis" by(erule_tac abc_append_first_steps_halt_eq) qed moreover have "\ stp. abc_steps_l (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything) ((?A [+] (?B [+] ?C)) @ ?D) stp = (0, xs @ (x - y) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" using len by(rule_tac loop_back, simp_all) moreover have "rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) = rec_exec (Pr n f g) (xs @ [x - y])" using less apply(cases "x - y", simp_all add: rec_exec_pr_Suc_simps) apply(rename_tac nat) by(subgoal_tac "nat = x - Suc y", simp, arith) ultimately show "?thesis" using code ft apply (auto simp add: abc_steps_add replicate_Suc_iff_anywhere) apply(rename_tac stp stpa) apply(rule_tac x = "stp + stpa" in exI) by (simp add: abc_steps_add replicate_Suc_iff_anywhere del: replicate_Suc) qed lemma abc_lm_s_simp0[simp]: "length xs = n \ abc_lm_s (xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ 0 # anything) (max (n + 3) (max fft gft)) 0 = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything" apply(simp add: abc_lm_s.simps) using list_update_length[of "xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n))" 0 anything 0] apply(auto simp: Suc_diff_Suc) apply(simp add: exp_suc[THEN sym] Suc_diff_Suc del: replicate_Suc) done lemma index_at_zero_elem[simp]: "(xs @ x # re # 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs))) @ 0 # anything) ! max (length xs + 3) (max fft gft) = 0" using nth_append_length[of "xs @ x # re # 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" 0 anything] by(simp) lemma pr_loop_correct: assumes less: "y \ x" and len: "length xs = n" and compile_g: "rec_ci g = (gap, gar, gft)" and termi_g: "\ y yanything. \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything\ gap \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything\)" shows "\\nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ y # anything\ ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)] \\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything\" using less proof(induct y) case 0 thus "?case" using len apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = 1 in exI) by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_comp.simps abc_shift.simps, simp add: abc_lm_v.simps) next case (Suc y) let ?ft = "max (n + 3) (max fft gft)" let ?C = "[Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" have ind: "y \ x \ \\nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything\ ?C \\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything\" by fact have less: "Suc y \ x" by fact have stp1: "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (?ft - Suc (Suc n)) @ Suc y # anything) ?C stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything)" using assms less by(rule_tac pr_loop, auto) then obtain stp1 where a: "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (?ft - Suc (Suc n)) @ Suc y # anything) ?C stp1 = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything)" .. moreover have "\ stp. abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) ?C stp = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything)" using ind less apply(auto simp: abc_Hoare_halt_def) apply(rename_tac na,case_tac "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) ?C na", rule_tac x = na in exI) by simp then obtain stp2 where b: "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) ?C stp2 = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything)" .. from a b show "?case" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add). qed lemma compile_pr_correct': assumes termi_g: "\ y yanything. \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything\ gap \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything\)" and termi_f: "terminate f xs" and f_ind: "\ anything. \\nl. nl = xs @ 0 \ (fft - far) @ anything\ fap \\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ anything\" and len: "length xs = n" and compile1: "rec_ci f = (fap, far, fft)" and compile2: "rec_ci g = (gap, gar, gft)" shows "\\nl. nl = xs @ x # 0 \ (max (n + 3) (max fft gft) - n) @ anything\ mv_box n (max (n + 3) (max fft gft)) [+] (fap [+] (mv_box n (Suc n) [+] ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]))) \\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything\" proof - let ?ft = "max (n+3) (max fft gft)" let ?A = "mv_box n ?ft" let ?B = "fap" let ?C = "mv_box n (Suc n)" let ?D = "[Dec ?ft (length gap + 7)]" let ?E = "gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]" let ?F = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" let ?P = "\nl. nl = xs @ x # 0 \ (?ft - n) @ anything" let ?S = "\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything" let ?Q1 = "\nl. nl = xs @ 0 \ (?ft - n) @ x # anything" show "\?P\ (?A [+] (?B [+] (?C [+] (?D [+] ?E @ ?F)))) \?S\" proof(rule_tac abc_Hoare_plus_halt) show "\?P\ ?A \?Q1\" using len by simp next let ?Q2 = "\nl. nl = xs @ rec_exec f xs # 0 \ (?ft - Suc n) @ x # anything" have a: "?ft \ fft" by arith have b: "far = n" using compile1 termi_f len by(drule_tac param_pattern, auto) have c: "fft > far" using compile1 by(simp add: footprint_ge) show "\?Q1\ (?B [+] (?C [+] (?D [+] ?E @ ?F))) \?S\" proof(rule_tac abc_Hoare_plus_halt) have "\\nl. nl = xs @ 0 \ (fft - far) @ 0\(?ft - fft) @ x # anything\ fap \\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ 0\(?ft - fft) @ x # anything\" by(rule_tac f_ind) moreover have "fft - far + ?ft - fft = ?ft - far" using a b c by arith moreover have "fft - Suc n + ?ft - fft = ?ft - Suc n" using a b c by arith ultimately show "\?Q1\ ?B \?Q2\" using b by(simp add: replicate_merge_anywhere) next let ?Q3 = "\ nl. nl = xs @ 0 # rec_exec f xs # 0\(?ft - Suc (Suc n)) @ x # anything" show "\?Q2\ (?C [+] (?D [+] ?E @ ?F)) \?S\" proof(rule_tac abc_Hoare_plus_halt) show "\?Q2\ (?C) \?Q3\" using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything"] using len by(auto) next show "\?Q3\ (?D [+] ?E @ ?F) \?S\" using pr_loop_correct[of x x xs n g gap gar gft f fft anything] assms by(simp add: rec_exec_pr_0_simps) qed qed qed qed lemma compile_pr_correct: assumes g_ind: "\y (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (xb - Suc xa) @ xc\))" and termi_f: "terminate f xs" and f_ind: "\ap arity fp anything. rec_ci f = (ap, arity, fp) \ \\nl. nl = xs @ 0 \ (fp - arity) @ anything\ ap \\nl. nl = xs @ rec_exec f xs # 0 \ (fp - Suc arity) @ anything\" and len: "length xs = n" and compile: "rec_ci (Pr n f g) = (ap, arity, fp)" shows "\\nl. nl = xs @ x # 0 \ (fp - arity) @ anything\ ap \\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (fp - Suc arity) @ anything\" proof(cases "rec_ci f", cases "rec_ci g") fix fap far fft gap gar gft assume h: "rec_ci f = (fap, far, fft)" "rec_ci g = (gap, gar, gft)" have g: "\y (\anything. \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything\ gap \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything\))" using g_ind h by(auto) hence termi_g: "\ y yanything. \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything\ gap \\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything\)" by auto have f_newind: "\ anything. \\nl. nl = xs @ 0 \ (fft - far) @ anything\ fap \\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ anything\" using h by(rule_tac f_ind, simp) show "?thesis" using termi_f termi_g h compile apply(simp add: rec_ci.simps abc_comp_commute, auto) using g_newind f_newind len by(rule_tac compile_pr_correct', simp_all) qed subsubsection \Correctness of compilation for constructor Mn\ fun mn_ind_inv :: "nat \ nat list \ nat \ nat \ nat \ nat list \ nat list \ bool" where "mn_ind_inv (as, lm') ss x rsx suf_lm lm = (if as = ss then lm' = lm @ x # rsx # suf_lm else if as = ss + 1 then \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx else if as = ss + 2 then \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm else False )" fun mn_stage1 :: "nat \ nat list \ nat \ nat \ nat" where "mn_stage1 (as, lm) ss n = (if as = 0 then 0 else if as = ss + 4 then 1 else if as = ss + 3 then 2 else if as = ss + 2 \ as = ss + 1 then 3 else if as = ss then 4 else 0 )" fun mn_stage2 :: "nat \ nat list \ nat \ nat \ nat" where "mn_stage2 (as, lm) ss n = (if as = ss + 1 \ as = ss + 2 then (lm ! (Suc n)) else 0)" fun mn_stage3 :: "nat \ nat list \ nat \ nat \ nat" where "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)" fun mn_measure :: "((nat \ nat list) \ nat \ nat) \ (nat \ nat \ nat)" where "mn_measure ((as, lm), ss, n) = (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n, mn_stage3 (as, lm) ss n)" definition mn_LE :: "(((nat \ nat list) \ nat \ nat) \ ((nat \ nat list) \ nat \ nat)) set" where "mn_LE \ (inv_image lex_triple mn_measure)" lemma wf_mn_le[intro]: "wf mn_LE" by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def) declare mn_ind_inv.simps[simp del] lemma put_in_tape_small_enough0[simp]: "0 < rsx \ \y. (xs @ x # rsx # anything)[Suc (length xs) := rsx - Suc 0] = xs @ x # y # anything \ y \ rsx" apply(rule_tac x = "rsx - 1" in exI) apply(simp add: list_update_append list_update.simps) done lemma put_in_tape_small_enough1[simp]: "\y \ rsx; 0 < y\ \ \ya. (xs @ x # y # anything)[Suc (length xs) := y - Suc 0] = xs @ x # ya # anything \ ya \ rsx" apply(rule_tac x = "y - 1" in exI) apply(simp add: list_update_append list_update.simps) done lemma abc_comp_null[simp]: "(A [+] B = []) = (A = [] \ B = [])" by(auto simp: abc_comp.simps abc_shift.simps) lemma rec_ci_not_null[simp]: "(rec_ci f \ ([], a, b))" proof(cases f) case (Cn x41 x42 x43) then show ?thesis by(cases "rec_ci x42", auto simp: mv_box.simps rec_ci.simps rec_ci_id.simps) next case (Pr x51 x52 x53) then show ?thesis apply(cases "rec_ci x52", cases "rec_ci x53") by (auto simp: mv_box.simps rec_ci.simps rec_ci_id.simps) next case (Mn x61 x62) then show ?thesis by(cases "rec_ci x62") (auto simp: rec_ci.simps rec_ci_id.simps) qed (auto simp: rec_ci_z_def rec_ci_s_def rec_ci.simps addition.simps rec_ci_id.simps) lemma mn_correct: assumes compile: "rec_ci rf = (fap, far, fft)" and ge: "0 < rsx" and len: "length xs = arity" and B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and f: "f = (\ stp. (abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp, (length fap), arity)) " and P: "P =(\ ((as, lm), ss, arity). as = 0)" and Q: "Q = (\ ((as, lm), ss, arity). mn_ind_inv (as, lm) (length fap) x rsx anything xs)" shows "\ stp. P (f stp) \ Q (f stp)" proof(rule_tac halt_lemma2) show "wf mn_LE" using wf_mn_le by simp next show "Q (f 0)" by(auto simp: Q f abc_steps_l.simps mn_ind_inv.simps) next have "fap \ []" using compile by auto thus "\ P (f 0)" by(auto simp: f P abc_steps_l.simps) next have "fap \ []" using compile by auto then have "\\ P (f stp); Q (f stp)\ \ Q (f (Suc stp)) \ (f (Suc stp), f stp) \ mn_LE" for stp using ge len apply(cases "(abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp)") apply(simp add: abc_step_red2 B f P Q) apply(auto split:if_splits simp add:abc_steps_l.simps mn_ind_inv.simps abc_steps_zero B abc_fetch.simps nth_append) by(auto simp: mn_LE_def lex_triple_def lex_pair_def abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps abc_lm_v.simps abc_lm_s.simps nth_append abc_fetch.simps split: if_splits) thus "\stp. \ P (f stp) \ Q (f stp) \ Q (f (Suc stp)) \ (f (Suc stp), f stp) \ mn_LE" by(auto) qed lemma abc_Hoare_haltE: "\\ nl. nl = lm1\ p \\ nl. nl = lm2\ \ \ stp. abc_steps_l (0, lm1) p stp = (length p, lm2)" by(auto simp:abc_Hoare_halt_def elim!: abc_holds_for.elims) lemma mn_loop: assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and far: "far = Suc arity" and ind: " (\xc. (\\nl. nl = xs @ x # 0 \ (fft - far) @ xc\ fap \\nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \ (fft - Suc far) @ xc\))" and exec_less: "rec_exec f (xs @ [x]) > 0" and compile: "rec_ci f = (fap, far, fft)" shows "\ stp > 0. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" proof - have "\ stp. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "\ stp. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "\\nl. nl = xs @ x # 0 \ (fft - far) @ 0\(ft - fft) @ anything\ fap \\nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \ (fft - Suc far) @ 0\(ft - fft) @ anything\" using ind by simp moreover have "fft > far" using compile by(erule_tac footprint_ge) ultimately show "?thesis" using ft far apply(drule_tac abc_Hoare_haltE) by(simp add: replicate_merge_anywhere max_absorb2) qed then obtain stp where "abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" .. thus "?thesis" by(erule_tac abc_append_first_steps_halt_eq) qed moreover have "\ stp > 0. abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 # 0 \ (ft - Suc (Suc arity)) @ anything)" using mn_correct[of f fap far fft "rec_exec f (xs @ [x])" xs arity B "(\stp. (abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp, length fap, arity))" x "0 \ (ft - Suc (Suc arity)) @ anything" "(\((as, lm), ss, arity). as = 0)" "(\((as, lm), ss, aritya). mn_ind_inv (as, lm) (length fap) x (rec_exec f (xs @ [x])) (0 \ (ft - Suc (Suc arity)) @ anything) xs) "] B compile exec_less len apply(subgoal_tac "fap \ []", auto) apply(rename_tac stp y) apply(rule_tac x = stp in exI, auto simp: mn_ind_inv.simps) by(case_tac "stp", simp_all add: abc_steps_l.simps) moreover have "fft > far" using compile by(erule_tac footprint_ge) ultimately show "?thesis" using ft far apply(auto) apply(rename_tac stp1 stp2) by(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add replicate_Suc[THEN sym] diff_Suc_Suc del: replicate_Suc) qed lemma mn_loop_correct': assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and ind_all: "\i\x. (\xc. (\\nl. nl = xs @ i # 0 \ (fft - far) @ xc\ fap \\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc\))" and exec_ge: "\ i\x. rec_exec f (xs @ [i]) > 0" and compile: "rec_ci f = (fap, far, fft)" and far: "far = Suc arity" shows "\ stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" using ind_all exec_ge proof(induct x) case 0 thus "?case" using assms by(rule_tac mn_loop, simp_all) next case (Suc x) have ind': "\\i\x. \xc. \\nl. nl = xs @ i # 0 \ (fft - far) @ xc\ fap \\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc\; \i\x. 0 < rec_exec f (xs @ [i])\ \ \stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" by fact have exec_ge: "\i\Suc x. 0 < rec_exec f (xs @ [i])" by fact have ind_all: "\i\Suc x. \xc. \\nl. nl = xs @ i # 0 \ (fft - far) @ xc\ fap \\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc\" by fact have ind: "\stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" using ind' exec_ge ind_all by simp have stp: "\ stp > 0. abc_steps_l (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc (Suc x) # 0 \ (ft - Suc arity) @ anything)" using ind_all exec_ge B ft len far compile by(rule_tac mn_loop, simp_all) from ind stp show "?case" apply(auto simp add: abc_steps_add) apply(rename_tac stp1 stp2) by(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add) qed lemma mn_loop_correct: assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and ind_all: "\i\x. (\xc. (\\nl. nl = xs @ i # 0 \ (fft - far) @ xc\ fap \\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc\))" and exec_ge: "\ i\x. rec_exec f (xs @ [i]) > 0" and compile: "rec_ci f = (fap, far, fft)" and far: "far = Suc arity" shows "\ stp. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" proof - have "\stp>x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" using assms by(rule_tac mn_loop_correct', simp_all) thus "?thesis" by(auto) qed lemma compile_mn_correct': assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and termi_f: "terminate f (xs @ [r])" and f_ind: "\anything. \\nl. nl = xs @ r # 0 \ (fft - far) @ anything\ fap \\nl. nl = xs @ r # 0 # 0 \ (fft - Suc far) @ anything\" and ind_all: "\i < r. (\xc. (\\nl. nl = xs @ i # 0 \ (fft - far) @ xc\ fap \\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc\))" and exec_less: "\ i 0" and exec: "rec_exec f (xs @ [r]) = 0" and compile: "rec_ci f = (fap, far, fft)" shows "\\nl. nl = xs @ 0 \ (max (Suc arity) fft - arity) @ anything\ fap @ B \\nl. nl = xs @ rec_exec (Mn arity f) xs # 0 \ (max (Suc arity) fft - Suc arity) @ anything\" proof - have a: "far = Suc arity" using len compile termi_f by(drule_tac param_pattern, auto) have b: "fft > far" using compile by(erule_tac footprint_ge) have "\ stp. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ r # 0 \ (ft - Suc arity) @ anything)" using assms a apply(cases r, rule_tac x = 0 in exI, simp add: abc_steps_l.simps, simp) by(rule_tac mn_loop_correct, auto) moreover have "\ stp. abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "\ stp. abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "\\nl. nl = xs @ r # 0 \ (fft - far) @ 0\(ft - fft) @ anything\ fap \\nl. nl = xs @ r # rec_exec f (xs @ [r]) # 0 \ (fft - Suc far) @ 0\(ft - fft) @ anything\" using f_ind exec by simp thus "?thesis" using ft a b apply(drule_tac abc_Hoare_haltE) by(simp add: replicate_merge_anywhere max_absorb2) qed then obtain stp where "abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" .. thus "?thesis" by(erule_tac abc_append_first_steps_halt_eq) qed moreover have "\ stp. abc_steps_l (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = (length fap + 5, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" using ft a b len B exec apply(rule_tac x = 1 in exI, auto) by(auto simp: abc_steps_l.simps B abc_step_l.simps abc_fetch.simps nth_append max_absorb2 abc_lm_v.simps abc_lm_s.simps) moreover have "rec_exec (Mn (length xs) f) xs = r" using exec exec_less apply(auto simp: rec_exec.simps Least_def) thm the_equality apply(rule_tac the_equality, auto) apply(metis exec_less less_not_refl3 linorder_not_less) by (metis le_neq_implies_less less_not_refl3) ultimately show "?thesis" using ft a b len B exec apply(auto simp: abc_Hoare_halt_def) apply(rename_tac stp1 stp2 stp3) apply(rule_tac x = "stp1 + stp2 + stp3" in exI) by(simp add: abc_steps_add replicate_Suc_iff_anywhere max_absorb2 Suc_diff_Suc del: replicate_Suc) qed lemma compile_mn_correct: assumes len: "length xs = n" and termi_f: "terminate f (xs @ [r])" and f_ind: "\ap arity fp anything. rec_ci f = (ap, arity, fp) \ \\nl. nl = xs @ r # 0 \ (fp - arity) @ anything\ ap \\nl. nl = xs @ r # 0 # 0 \ (fp - Suc arity) @ anything\" and exec: "rec_exec f (xs @ [r]) = 0" and ind_all: "\i (\x xa xb. rec_ci f = (x, xa, xb) \ (\xc. \\nl. nl = xs @ i # 0 \ (xb - xa) @ xc\ x \\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (xb - Suc xa) @ xc\)) \ 0 < rec_exec f (xs @ [i])" and compile: "rec_ci (Mn n f) = (ap, arity, fp)" shows "\\nl. nl = xs @ 0 \ (fp - arity) @ anything\ ap \\nl. nl = xs @ rec_exec (Mn n f) xs # 0 \ (fp - Suc arity) @ anything\" proof(cases "rec_ci f") fix fap far fft assume h: "rec_ci f = (fap, far, fft)" hence f_newind: "\anything. \\nl. nl = xs @ r # 0 \ (fft - far) @ anything\ fap \\nl. nl = xs @ r # 0 # 0 \ (fft - Suc far) @ anything\" by(rule_tac f_ind, simp) have newind_all: "\i < r. (\xc. (\\nl. nl = xs @ i # 0 \ (fft - far) @ xc\ fap \\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc\))" using ind_all h by(auto) have all_less: "\ i 0" using ind_all by auto show "?thesis" using h compile f_newind newind_all all_less len termi_f exec apply(auto simp: rec_ci.simps) by(rule_tac compile_mn_correct', auto) qed subsubsection \Correctness of entire compilation process rec\_ci\ lemma recursive_compile_correct: "\terminate recf args; rec_ci recf = (ap, arity, fp)\ \ \\ nl. nl = args @ 0\(fp - arity) @ anything\ ap \\ nl. nl = args@ rec_exec recf args # 0\(fp - Suc arity) @ anything\" apply(induct arbitrary: ap arity fp anything rule: terminate.induct) apply(simp_all add: compile_s_correct compile_z_correct compile_id_correct compile_cn_correct compile_pr_correct compile_mn_correct) done definition dummy_abc :: "nat \ abc_inst list" where "dummy_abc k = [Inc k, Dec k 0, Goto 3]" definition abc_list_crsp:: "nat list \ nat list \ bool" where "abc_list_crsp xs ys = (\ n. xs = ys @ 0\n \ ys = xs @ 0\n)" lemma abc_list_crsp_simp1[intro]: "abc_list_crsp (lm @ 0\m) lm" by(auto simp: abc_list_crsp_def) lemma abc_list_crsp_lm_v: "abc_list_crsp lma lmb \ abc_lm_v lma n = abc_lm_v lmb n" by(auto simp: abc_list_crsp_def abc_lm_v.simps nth_append) lemma abc_list_crsp_elim: "\abc_list_crsp lma lmb; \ n. lma = lmb @ 0\n \ lmb = lma @ 0\n \ P \ \ P" by(auto simp: abc_list_crsp_def) lemma abc_list_crsp_simp[simp]: "\abc_list_crsp lma lmb; m < length lma; m < length lmb\ \ abc_list_crsp (lma[m := n]) (lmb[m := n])" by(auto simp: abc_list_crsp_def list_update_append) lemma abc_list_crsp_simp2[simp]: "\abc_list_crsp lma lmb; m < length lma; \ m < length lmb\ \ abc_list_crsp (lma[m := n]) (lmb @ 0 \ (m - length lmb) @ [n])" apply(auto simp: abc_list_crsp_def list_update_append) apply(rename_tac N) apply(rule_tac x = "N + length lmb - Suc m" in exI) apply(rule_tac disjI1) apply(simp add: upd_conv_take_nth_drop min_absorb1) done lemma abc_list_crsp_simp3[simp]: "\abc_list_crsp lma lmb; \ m < length lma; m < length lmb\ \ abc_list_crsp (lma @ 0 \ (m - length lma) @ [n]) (lmb[m := n])" apply(auto simp: abc_list_crsp_def list_update_append) apply(rename_tac N) apply(rule_tac x = "N + length lma - Suc m" in exI) apply(rule_tac disjI2) apply(simp add: upd_conv_take_nth_drop min_absorb1) done lemma abc_list_crsp_simp4[simp]: "\abc_list_crsp lma lmb; \ m < length lma; \ m < length lmb\ \ abc_list_crsp (lma @ 0 \ (m - length lma) @ [n]) (lmb @ 0 \ (m - length lmb) @ [n])" by(auto simp: abc_list_crsp_def list_update_append replicate_merge_anywhere) lemma abc_list_crsp_lm_s: "abc_list_crsp lma lmb \ abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)" by(auto simp: abc_lm_s.simps) lemma abc_list_crsp_step: "\abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); abc_step_l (aa, lmb) i = (a', lmb')\ \ a' = a \ abc_list_crsp lma' lmb'" apply(cases i, auto simp: abc_step_l.simps abc_list_crsp_lm_s abc_list_crsp_lm_v split: abc_inst.splits if_splits) done lemma abc_list_crsp_steps: "\abc_steps_l (0, lm @ 0\m) aprog stp = (a, lm'); aprog \ []\ \ \ lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ abc_list_crsp lm' lma" proof(induct stp arbitrary: a lm') case (Suc stp) then show ?case using [[simproc del: defined_all]] apply(cases "abc_steps_l (0, lm @ 0\m) aprog stp", simp add: abc_step_red) proof - fix stp a lm' aa b assume ind: "\a lm'. aa = a \ b = lm' \ \lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ abc_list_crsp lm' lma" and h: "abc_step_l (aa, b) (abc_fetch aa aprog) = (a, lm')" "abc_steps_l (0, lm @ 0\m) aprog stp = (aa, b)" "aprog \ []" have "\lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \ abc_list_crsp b lma" apply(rule_tac ind, simp) done from this obtain lma where g2: "abc_steps_l (0, lm) aprog stp = (aa, lma) \ abc_list_crsp b lma" .. hence g3: "abc_steps_l (0, lm) aprog (Suc stp) = abc_step_l (aa, lma) (abc_fetch aa aprog)" apply(rule_tac abc_step_red, simp) done show "\lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \ abc_list_crsp lm' lma" using g2 g3 h apply(auto) apply(cases "abc_step_l (aa, b) (abc_fetch aa aprog)", cases "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp) apply(rule_tac abc_list_crsp_step, auto) done qed qed (force simp add: abc_steps_l.simps) lemma list_crsp_simp2: "abc_list_crsp (lm1 @ 0\n) lm2 \ abc_list_crsp lm1 lm2" proof(induct n) case 0 thus "?case" by(auto simp: abc_list_crsp_def) next case (Suc n) have ind: "abc_list_crsp (lm1 @ 0 \ n) lm2 \ abc_list_crsp lm1 lm2" by fact have h: "abc_list_crsp (lm1 @ 0 \ Suc n) lm2" by fact then have "abc_list_crsp (lm1 @ 0 \ n) lm2" apply(auto simp only: exp_suc abc_list_crsp_def ) apply (metis Suc_pred append_eq_append_conv append_eq_append_conv2 butlast_append butlast_snoc length_replicate list.distinct(1) neq0_conv replicate_Suc replicate_Suc_iff_anywhere replicate_app_Cons_same replicate_empty self_append_conv self_append_conv2) apply (auto,metis replicate_Suc) . thus "?case" using ind by auto qed lemma recursive_compile_correct_norm': "\rec_ci f = (ap, arity, ft); terminate f args\ \ \ stp rl. (abc_steps_l (0, args) ap stp) = (length ap, rl) \ abc_list_crsp (args @ [rec_exec f args]) rl" using recursive_compile_correct[of f args ap arity ft "[]"] apply(auto simp: abc_Hoare_halt_def) apply(rename_tac n) apply(rule_tac x = n in exI) apply(case_tac "abc_steps_l (0, args @ 0 \ (ft - arity)) ap n", auto) apply(drule_tac abc_list_crsp_steps, auto) apply(rule_tac list_crsp_simp2, auto) done lemma find_exponent_rec_exec[simp]: assumes a: "args @ [rec_exec f args] = lm @ 0 \ n" and b: "length args < length lm" shows "\m. lm = args @ rec_exec f args # 0 \ m" using assms apply(cases n, simp) apply(rule_tac x = 0 in exI, simp) apply(drule_tac length_equal, simp) done lemma find_exponent_complex[simp]: "\args @ [rec_exec f args] = lm @ 0 \ n; \ length args < length lm\ \ \m. (lm @ 0 \ (length args - length lm) @ [Suc 0])[length args := 0] = args @ rec_exec f args # 0 \ m" apply(cases n, simp_all add: exp_suc list_update_append list_update.simps del: replicate_Suc) apply(subgoal_tac "length args = Suc (length lm)", simp) apply(rule_tac x = "Suc (Suc 0)" in exI, simp) apply(drule_tac length_equal, simp, auto) done lemma compile_append_dummy_correct: assumes compile: "rec_ci f = (ap, ary, fp)" and termi: "terminate f args" shows "\\ nl. nl = args\ (ap [+] dummy_abc (length args)) \\ nl. (\ m. nl = args @ rec_exec f args # 0\m)\" proof(rule_tac abc_Hoare_plus_halt) show "\\nl. nl = args\ ap \\ nl. abc_list_crsp (args @ [rec_exec f args]) nl\" using compile termi recursive_compile_correct_norm'[of f ap ary fp args] apply(auto simp: abc_Hoare_halt_def) by (metis abc_final.simps abc_holds_for.simps) next show "\abc_list_crsp (args @ [rec_exec f args])\ dummy_abc (length args) \\nl. \m. nl = args @ rec_exec f args # 0 \ m\" apply(auto simp: dummy_abc_def abc_Hoare_halt_def) apply(rule_tac x = 3 in exI) by(force simp: abc_steps_l.simps abc_list_crsp_def abc_step_l.simps numeral_3_eq_3 abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps) qed lemma cn_merge_gs_split: "\i < length gs; rec_ci (gs!i) = (ga, gb, gc)\ \ cn_merge_gs (map rec_ci gs) p = cn_merge_gs (map rec_ci (take i gs)) p [+] (ga [+] mv_box gb (p + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)" proof(induct i arbitrary: gs p) case 0 then show ?case by(cases gs; simp) next case (Suc i) then show ?case by(cases gs, simp, cases "rec_ci (hd gs)", simp add: abc_comp_commute[THEN sym]) qed lemma cn_unhalt_case: assumes compile1: "rec_ci (Cn n f gs) = (ap, ar, ft) \ length args = ar" and g: "i < length gs" and compile2: "rec_ci (gs!i) = (gap, gar, gft) \ gar = length args" and g_unhalt: "\ anything. \\ nl. nl = args @ 0\(gft - gar) @ anything\ gap \" and g_ind: "\ apj arj ftj j anything. \j < i; rec_ci (gs!j) = (apj, arj, ftj)\ \ \\ nl. nl = args @ 0\(ftj - arj) @ anything\ apj \\ nl. nl = args @ rec_exec (gs!j) args # 0\(ftj - Suc arj) @ anything\" and all_termi: "\ j\ nl. nl = args @ 0\(ft - ar) @ anything\ ap \" using compile1 apply(cases "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) proof(rule_tac abc_Hoare_plus_unhalt1) fix fap far fft let ?ft = "max (Suc (length args)) (Max (insert fft ((\(aprog, p, n). n) ` rec_ci ` set gs)))" let ?Q = "\nl. nl = args @ 0\ (?ft - length args) @ map (\i. rec_exec i args) (take i gs) @ 0\(length gs - i) @ 0\ Suc (length args) @ anything" have "cn_merge_gs (map rec_ci gs) ?ft = cn_merge_gs (map rec_ci (take i gs)) ?ft [+] (gap [+] mv_box gar (?ft + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?ft + Suc i)" using g compile2 cn_merge_gs_split by simp thus "\\nl. nl = args @ 0 # 0 \ (?ft + length gs) @ anything\ (cn_merge_gs (map rec_ci gs) ?ft) \" proof(simp, rule_tac abc_Hoare_plus_unhalt1, rule_tac abc_Hoare_plus_unhalt2, rule_tac abc_Hoare_plus_unhalt1) let ?Q_tmp = "\nl. nl = args @ 0\ (gft - gar) @ 0\(?ft - (length args) - (gft -gar)) @ map (\i. rec_exec i args) (take i gs) @ 0\(length gs - i) @ 0\ Suc (length args) @ anything" have a: "\?Q_tmp\ gap \" using g_unhalt[of "0 \ (?ft - (length args) - (gft - gar)) @ map (\i. rec_exec i args) (take i gs) @ 0 \ (length gs - i) @ 0 \ Suc (length args) @ anything"] by simp moreover have "?ft \ gft" using g compile2 apply(rule_tac max.coboundedI2, rule_tac Max_ge, simp, rule_tac insertI2) apply(rule_tac x = "rec_ci (gs ! i)" in image_eqI, simp) by(rule_tac x = "gs!i" in image_eqI, simp, simp) then have b:"?Q_tmp = ?Q" using compile2 apply(rule_tac arg_cong) by(simp add: replicate_merge_anywhere) thus "\?Q\ gap \" using a by simp next show "\\nl. nl = args @ 0 # 0 \ (?ft + length gs) @ anything\ cn_merge_gs (map rec_ci (take i gs)) ?ft \\nl. nl = args @ 0 \ (?ft - length args) @ map (\i. rec_exec i args) (take i gs) @ 0 \ (length gs - i) @ 0 \ Suc (length args) @ anything\" using all_termi by(rule_tac compile_cn_gs_correct', auto simp: set_conv_nth intro:g_ind) qed qed lemma mn_unhalt_case': assumes compile: "rec_ci f = (a, b, c)" and all_termi: "\i. terminate f (args @ [i]) \ 0 < rec_exec f (args @ [i])" and B: "B = [Dec (Suc (length args)) (length a + 5), Dec (Suc (length args)) (length a + 3), Goto (Suc (length a)), Inc (length args), Goto 0]" shows "\\nl. nl = args @ 0 \ (max (Suc (length args)) c - length args) @ anything\ a @ B \" proof(rule_tac abc_Hoare_unhaltI, auto) fix n have a: "b = Suc (length args)" using all_termi compile apply(erule_tac x = 0 in allE) by(auto, drule_tac param_pattern,auto) moreover have b: "c > b" using compile by(elim footprint_ge) ultimately have c: "max (Suc (length args)) c = c" by arith have "\ stp > n. abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) stp = (0, args @ Suc n # 0\(c - Suc (length args)) @ anything)" using assms a b c proof(rule_tac mn_loop_correct', auto) fix i xc show "\\nl. nl = args @ i # 0 \ (c - Suc (length args)) @ xc\ a \\nl. nl = args @ i # rec_exec f (args @ [i]) # 0 \ (c - Suc (Suc (length args))) @ xc\" using all_termi recursive_compile_correct[of f "args @ [i]" a b c xc] compile a by(simp) qed then obtain stp where d: "stp > n \ abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) stp = (0, args @ Suc n # 0\(c - Suc (length args)) @ anything)" .. then obtain d where e: "stp = n + Suc d" by (metis add_Suc_right less_iff_Suc_add) obtain s nl where f: "abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) n = (s, nl)" by (metis prod.exhaust) have g: "s < length (a @ B)" using d e f apply(rule_tac classical, simp only: abc_steps_add) by(simp add: halt_steps2 leI) from f g show "abc_notfinal (abc_steps_l (0, args @ 0 \ (max (Suc (length args)) c - length args) @ anything) (a @ B) n) (a @ B)" using c b a by(simp add: replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc) qed lemma mn_unhalt_case: assumes compile: "rec_ci (Mn n f) = (ap, ar, ft) \ length args = ar" and all_term: "\ i. terminate f (args @ [i]) \ rec_exec f (args @ [i]) > 0" shows "\ (\ nl. nl = args @ 0\(ft - ar) @ anything) \ ap \ " using assms apply(cases "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) by(rule_tac mn_unhalt_case', simp_all) section \Compilers composed: Compiling Recursive Functions into Turing Machines\ fun tm_of_rec :: "recf \ instr list" where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in let tp = tm_of (ap [+] dummy_abc k) in tp @ (shift (mopup_n_tm k) (length tp div 2)))" lemma recursive_compile_to_tm_correct1: assumes compile: "rec_ci recf = (ap, ary, fp)" and termi: " terminate recf args" and tp: "tp = tm_of (ap [+] dummy_abc (length args))" shows "\ stp m l. steps0 (Suc 0, Bk # Bk # ires, @ Bk\rn) (tp @ shift (mopup_n_tm (length args)) (length tp div 2)) stp = (0, Bk\m @ Bk # Bk # ires, Oc\Suc (rec_exec recf args) @ Bk\l)" proof - have "\\nl. nl = args\ ap [+] dummy_abc (length args) \\nl. \m. nl = args @ rec_exec recf args # 0 \ m\" using compile termi compile by(rule_tac compile_append_dummy_correct, auto) then obtain stp m where h: "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) stp = (length (ap [+] dummy_abc (length args)), args @ rec_exec recf args # 0\m) " apply(simp add: abc_Hoare_halt_def, auto) apply(rename_tac n) by(case_tac "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) n", auto) thus "?thesis" using assms tp compile_correct_halt[OF refl refl _ h _ _ refl] by(auto simp: crsp.simps start_of.simps abc_lm_v.simps) qed lemma recursive_compile_to_tm_correct2: assumes termi: " terminate recf args" shows "\ stp m l. steps0 (Suc 0, [Bk, Bk], ) (tm_of_rec recf) stp = (0, Bk\Suc (Suc m), Oc\Suc (rec_exec recf args) @ Bk\l)" proof(cases "rec_ci recf", simp ) fix ap ar fp assume "rec_ci recf = (ap, ar, fp)" thus "\stp m l. steps0 (Suc 0, [Bk, Bk], ) (tm_of (ap [+] dummy_abc ar) @ shift (mopup_n_tm ar) (sum_list (layout_of (ap [+] dummy_abc ar)))) stp = (0, Bk # Bk # Bk \ m, Oc # Oc \ rec_exec recf args @ Bk \ l)" using recursive_compile_to_tm_correct1[of recf ap ar fp args "tm_of (ap [+] dummy_abc (length args))" "[]" 0] assms param_pattern[of recf args ap ar fp] by(simp add: replicate_Suc[THEN sym] replicate_Suc_iff_anywhere del: replicate_Suc, simp add: exp_suc del: replicate_Suc) qed lemma recursive_compile_to_tm_correct3: assumes termi: "terminate recf args" shows "\\ tp. tp =([Bk, Bk], )\ (tm_of_rec recf) \\ tp. \ k l. tp = (Bk\ k, @ Bk \ l)\" using recursive_compile_to_tm_correct2[OF assms] apply(auto simp add: Hoare_halt_def ) apply(rename_tac stp M l) apply(rule_tac x = stp in exI) apply(auto simp add: tape_of_nat_def) apply(rule_tac x = "Suc (Suc M)" in exI) apply(simp) done subsection \Appending the mopup TM\ lemma list_all_suc_many[simp]: "list_all (\(acn, s). s \ Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \ list_all (\(acn, s). s \ Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs" proof(induct xs) case (Cons a xs) then show ?case by(cases a, simp) qed simp lemma shift_append: "shift (xs @ ys) n = shift xs n @ shift ys n" apply(simp add: shift.simps) done lemma length_shift_mopup[simp]: "length (shift (mopup_n_tm n) ss) = 4 * n + 12" apply(auto simp: mopup_n_tm.simps shift_append mopup_b_def) done lemma length_tm_even[intro]: "length (tm_of ap) mod 2 = 0" apply(simp add: tm_of.simps) done lemma tms_of_at_index[simp]: "k < length ap \ tms_of ap ! k = ci (layout_of ap) (start_of (layout_of ap) k) (ap ! k)" apply(simp add: tms_of.simps tpairs_of.simps) done lemma start_of_suc_inc: "\k < length ap; ap ! k = Inc n\ \ start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2 * n + 9" apply(rule_tac start_of_Suc1, auto simp: abc_fetch.simps) done lemma start_of_suc_dec: "\k < length ap; ap ! k = (Dec n e)\ \ start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2 * n + 16" apply(rule_tac start_of_Suc2, auto simp: abc_fetch.simps) done lemma inc_state_all_le: "\k < length ap; ap ! k = Inc n; (a, b) \ set (shift (shift tinc_b (2 * n)) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") apply(subgoal_tac "start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap) ") apply(arith) apply(cases "Suc k = length ap", simp) apply(rule_tac start_of_less, simp) apply(auto simp: tinc_b_def shift.simps start_of_suc_inc length_of.simps ) done lemma findnth_le[elim]: "(a, b) \ set (shift (findnth n) (start_of (layout_of ap) k - Suc 0)) \ b \ Suc (start_of (layout_of ap) k + 2 * n)" apply(induct n, force simp add: shift.simps) apply(simp add: shift_append, auto) apply(auto simp: shift.simps) done lemma findnth_state_all_le1: "\k < length ap; ap ! k = Inc n; (a, b) \ set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") apply(subgoal_tac "start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap) ") apply(arith) apply(cases "Suc k = length ap", simp) apply(rule_tac start_of_less, simp) apply(subgoal_tac "b \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) (Suc k)", auto) apply(auto simp: tinc_b_def shift.simps length_of.simps start_of_suc_inc) done lemma start_of_eq: "length ap < as \ start_of (layout_of ap) as = start_of (layout_of ap) (length ap)" proof(induct as) case (Suc as) then show ?case apply(cases "length ap < as", simp add: start_of.simps) apply(subgoal_tac "as = length ap") apply(simp add: start_of.simps) apply arith done qed simp lemma start_of_all_le: "start_of (layout_of ap) as \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "as > length ap \ as = length ap \ as < length ap", auto simp: start_of_eq start_of_less) done lemma findnth_state_all_le2: "\k < length ap; ap ! k = Dec n e; (a, b) \ set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap)", auto) apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16", simp) apply(simp add: start_of_suc_dec) apply(rule_tac start_of_all_le) done lemma dec_state_all_le[simp]: "\k < length ap; ap ! k = Dec n e; (a, b) \ set (shift (shift tdec_b (2 * n)) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \ start_of (layout_of ap) (length ap) \ start_of (layout_of ap) k > 0") prefer 2 apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16 \ start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap)") apply(simp, rule_tac conjI) apply(simp add: start_of_suc_dec) apply(rule_tac start_of_all_le) apply(auto simp: tdec_b_def shift.simps) done lemma tms_any_less: "\k < length ap; (a, b) \ set (tms_of ap ! k)\ \ b \ start_of (layout_of ap) (length ap)" apply(cases "ap!k", auto simp: tms_of.simps tpairs_of.simps ci.simps shift_append adjust.simps) apply(erule_tac findnth_state_all_le1, simp_all) apply(erule_tac inc_state_all_le, simp_all) apply(erule_tac findnth_state_all_le2, simp_all) apply(rule_tac start_of_all_le) apply(rule_tac start_of_all_le) done lemma concat_in: "i < length (concat xs) \ \k < length xs. concat xs ! i \ set (xs ! k)" proof(induct xs rule: rev_induct) case (snoc x xs) then show ?case apply(cases "i < length (concat xs)", simp) apply(erule_tac exE, rule_tac x = k in exI) apply(simp add: nth_append) apply(rule_tac x = "length xs" in exI, simp) apply(simp add: nth_append) done qed auto declare length_concat[simp] lemma in_tms: "i < length (tm_of ap) \ \ k < length ap. (tm_of ap ! i) \ set (tms_of ap ! k)" apply(simp only: tm_of.simps) using concat_in[of i "tms_of ap"] apply(auto) done lemma all_le_start_of: "list_all (\(acn, s). s \ start_of (layout_of ap) (length ap)) (tm_of ap)" apply(simp only: list_all_length) apply(rule_tac allI, rule_tac impI) apply(drule_tac in_tms, auto elim: tms_any_less) done lemma length_ci: "\k < length ap; length (ci ly y (ap ! k)) = 2 * qa\ \ layout_of ap ! k = qa" apply(cases "ap ! k") apply(auto simp: layout_of.simps ci.simps length_of.simps tinc_b_def tdec_b_def length_findnth adjust.simps) done lemma ci_even[intro]: "length (ci ly y i) mod 2 = 0" apply(cases i, auto simp: ci.simps length_findnth tinc_b_def adjust.simps tdec_b_def) done lemma sum_list_ci_even[intro]: "sum_list (map (length \ (\(x, y). ci ly x y)) zs) mod 2 = 0" proof(induct zs rule: rev_induct) case (snoc x xs) then show ?case apply(cases x, simp) apply(subgoal_tac "length (ci ly (fst x) (snd x)) mod 2 = 0") apply(auto) done qed (simp) lemma zip_pre: "(length ys) \ length ap \ zip ys ap = zip ys (take (length ys) (ap::'a list))" proof(induct ys arbitrary: ap) case (Cons a ys) from Cons(2) have z:"ap = aa # list \ zip (a # ys) ap = zip (a # ys) (take (length (a # ys)) ap)" for aa list using Cons(1)[of list] by simp thus ?case by (cases ap;simp) qed simp lemma length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap) div 2)" using tpa_states[of "tm_of ap" "length ap" ap] by(simp add: tm_of.simps) lemma list_all_add_6E[elim]: "list_all (\(acn, s). s \ Suc q) xs \ list_all (\(acn, s). s \ q + (2 * n + 6)) xs" by(auto simp: list_all_length) lemma mopup_b_12[simp]: "length mopup_b = 12" by(simp add: mopup_b_def) lemma mp_up_all_le: "list_all (\(acn, s). s \ q + (2 * n + 6)) [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), (L, 5 + 2 * n + q), (WB, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q), (WB, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))), (WB, Suc (Suc (Suc (2 * n + q)))), (L, 5 + 2 * n + q), (L, 6 + 2 * n + q), (R, 0), (L, 6 + 2 * n + q)]" by(auto) lemma mopup_le6[simp]: "(a, b) \ set (mopup_a n) \ b \ 2 * n + 6" by(induct n, auto ) lemma shift_le2[simp]: "(a, b) \ set (shift (mopup_n_tm n) x) \ b \ (2 * x + length (mopup_n_tm n)) div 2" apply(auto simp: mopup_n_tm.simps shift_append shift.simps) apply(auto simp: mopup_b_def) done lemma mopup_ge2[intro]: " 2 \ x + length (mopup_n_tm n)" apply(simp add: mopup_n_tm.simps) done lemma mopup_even[intro]: " (2 * x + length (mopup_n_tm n)) mod 2 = 0" by(auto simp: mopup_n_tm.simps) lemma mopup_div_2[simp]: "b \ Suc x \ b \ (2 * x + length (mopup_n_tm n)) div 2" by(auto simp: mopup_n_tm.simps) subsection \A Turing Machine compiled from an Abacus program with mopup code appended is composable\ lemma composable_tm_from_abacus: assumes "tp = tm_of ap" shows "composable_tm0 (tp @ shift (mopup_n_tm n) (length tp div 2))" proof - have "is_even (length (mopup_n_tm n))" for n using composable_tm.simps by blast moreover have "(aa, ba) \ set (mopup_n_tm n) \ ba \ length (mopup_n_tm n) div 2" for aa ba by (metis (no_types, lifting) add_cancel_left_right case_prodD composable_tm.simps composable_mopup_n_tm) moreover have "(\x\set (tm_of ap). case x of (acn, s::nat) \ s \ Suc (sum_list (layout_of ap))) \ (a, b) \ set (tm_of ap) \ b \ sum_list (layout_of ap) + length (mopup_n_tm n) div 2" for a b and s::nat by (metis (no_types, lifting) add_Suc add_cancel_left_right case_prodD div_mult_mod_eq le_SucE mult_2_right not_numeral_le_zero composable_tm.simps trans_le_add1 composable_mopup_n_tm) ultimately show ?thesis unfolding assms using length_start_of_tm[of ap] all_le_start_of[of ap] composable_tm.simps by(auto simp: List.list_all_iff shift.simps) qed subsection \A Turing Machine compiled from a recursive function is composable\ lemma composable_tm_from_recf: assumes compile: "tp = tm_of_rec recf" shows "composable_tm0 tp" proof - obtain a b c where "rec_ci recf = (a, b, c)" by (metis prod_cases3) thus "?thesis" using compile using composable_tm_from_abacus[of "tm_of (a [+] dummy_abc b)" "(a [+] dummy_abc b)" b] by simp qed end diff --git a/thys/Universal_Turing_Machine/StrongCopyTM.thy b/thys/Universal_Turing_Machine/StrongCopyTM.thy --- a/thys/Universal_Turing_Machine/StrongCopyTM.thy +++ b/thys/Universal_Turing_Machine/StrongCopyTM.thy @@ -1,5696 +1,5696 @@ (* Title: thys/StrongCopyTM.thy Author: Franz Regensburger (FABR) 02/2022 *) subsubsection \A Turing machine that duplicates its input iff the input is a single numeral\ theory StrongCopyTM imports WeakCopyTM begin text \If we run \tm_strong_copy\ on a single numeral, it behaves like the original weak version \tm_weak_copy\. However, if we run the strong machine on an empty list, the result is an empty list. If we run the machine on a list with more than two numerals, this strong variant will just return the first numeral of the list (a singleton list). Thus, the result will be a list of two numerals only if we run it on a singleton list. This is exactly the property, we need for the reduction of problem \K1\ to problem \H1\. \ (* ---------- tm_skip_first_arg ---------- *) (* let tm_skip_first_arg = from_raw [ (L,0),(R,2), -- 1 on Bk stop and delegate to tm_erase_right_then_dblBk_left, on Oc investigate further (R,3),(R,2), -- 2 'go right until Bk': on Bk check for further args, on Oc continue loop 'go right until Bk' (L,4),(WO,0), -- 3 on 2nd Bk go for OK path, on Oc delegate to tm_erase_right_then_dblBk_left -- (L,5),(L,5), -- 4 skip 1st Bk -- (R,0),(L,5) -- 5 'go left until Bk': on 2nd Bk stop, on Oc continue loop 'go left until Bk' ] *) definition tm_skip_first_arg :: "instr list" where "tm_skip_first_arg \ [ (L,0),(R,2), (R,3),(R,2), (L,4),(WO,0), (L,5),(L,5), (R,0),(L,5) ]" (* Prove partial correctness and termination for tm_skip_first_arg *) (* ERROR case: length nl = 0 *) lemma tm_skip_first_arg_correct_Nil: "\\tap. tap = ([], [])\ tm_skip_first_arg \\tap. tap = ([], [Bk]) \" proof - have "steps0 (1, [], []) tm_skip_first_arg 1 = (0::nat, [], [Bk])" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_skip_first_arg_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq) qed corollary tm_skip_first_arg_correct_Nil': "length nl = 0 \ \\tap. tap = ([], )\ tm_skip_first_arg \\tap. tap = ([], [Bk]) \" using tm_skip_first_arg_correct_Nil by (simp add: tm_skip_first_arg_correct_Nil ) (* OK cases: length nl = 1 *) (* ctxrunTM tm_skip_first_arg (1, [], [Oc]) 0: [] _1_ [Oc] => 1: [Oc] _2_ [] => 2: [Oc,Bk] _3_ [] => 3: [Oc] _4_ [Bk] => 4: [] _5_ [Oc,Bk] => 5: [] _5_ [Bk,Oc,Bk] => 6: [Bk] _0_ [Oc,Bk] ctxrunTM tm_skip_first_arg (1, [], [Oc, Oc, Oc, Oc]) 0: [] _1_ [Oc,Oc,Oc,Oc] => 1: [Oc] _2_ [Oc,Oc,Oc] => 2: [Oc,Oc] _2_ [Oc,Oc] => 3: [Oc,Oc,Oc] _2_ [Oc] => 4: [Oc,Oc,Oc,Oc] _2_ [] => 5: [Oc,Oc,Oc,Oc,Bk] _3_ [] => 6: [Oc,Oc,Oc,Oc] _4_ [Bk] => 7: [Oc,Oc,Oc] _5_ [Oc,Bk] => 8: [Oc,Oc] _5_ [Oc,Oc,Bk] => 9: [Oc] _5_ [Oc,Oc,Oc,Bk] => 10: [] _5_ [Oc,Oc,Oc,Oc,Bk] => 11: [] _5_ [Bk,Oc,Oc,Oc,Oc,Bk] => 12: [Bk] _0_ [Oc,Oc,Oc,Oc,Bk] *) fun inv_tm_skip_first_arg_len_eq_1_s0 :: "nat \ tape \ bool" and inv_tm_skip_first_arg_len_eq_1_s1 :: "nat \ tape \ bool" and inv_tm_skip_first_arg_len_eq_1_s2 :: "nat \ tape \ bool" and inv_tm_skip_first_arg_len_eq_1_s3 :: "nat \ tape \ bool" and inv_tm_skip_first_arg_len_eq_1_s4 :: "nat \ tape \ bool" and inv_tm_skip_first_arg_len_eq_1_s5 :: "nat \ tape \ bool" where "inv_tm_skip_first_arg_len_eq_1_s0 n (l, r) = ( l = [Bk] \ r = Oc \ (Suc n) @ [Bk])" | "inv_tm_skip_first_arg_len_eq_1_s1 n (l, r) = ( l = [] \ r = Oc \ Suc n)" | "inv_tm_skip_first_arg_len_eq_1_s2 n (l, r) = (\n1 n2. l = Oc \ (Suc n1) \ r = Oc \ n2 \ Suc n1 + n2 = Suc n)" | "inv_tm_skip_first_arg_len_eq_1_s3 n (l, r) = ( l = Bk # Oc \ (Suc n) \ r = [])" | "inv_tm_skip_first_arg_len_eq_1_s4 n (l, r) = ( l = Oc \ (Suc n) \ r = [Bk])" | "inv_tm_skip_first_arg_len_eq_1_s5 n (l, r) = (\n1 n2. (l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) \ (l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) )" fun inv_tm_skip_first_arg_len_eq_1 :: "nat \ config \ bool" where "inv_tm_skip_first_arg_len_eq_1 n (s, tap) = (if s = 0 then inv_tm_skip_first_arg_len_eq_1_s0 n tap else if s = 1 then inv_tm_skip_first_arg_len_eq_1_s1 n tap else if s = 2 then inv_tm_skip_first_arg_len_eq_1_s2 n tap else if s = 3 then inv_tm_skip_first_arg_len_eq_1_s3 n tap else if s = 4 then inv_tm_skip_first_arg_len_eq_1_s4 n tap else if s = 5 then inv_tm_skip_first_arg_len_eq_1_s5 n tap else False)" lemma tm_skip_first_arg_len_eq_1_cases: fixes s::nat assumes "inv_tm_skip_first_arg_len_eq_1 n (s,l,r)" and "s=0 \ P" and "s=1 \ P" and "s=2 \ P" and "s=3 \ P" and "s=4 \ P" and "s=5 \ P" shows "P" proof - have "s < 6" proof (rule ccontr) assume "\ s < 6" with \inv_tm_skip_first_arg_len_eq_1 n (s,l,r)\ show False by auto qed then have "s = 0 \ s = 1 \ s = 2 \ s = 3 \ s = 4 \ s = 5" by auto with assms show ?thesis by auto qed lemma inv_tm_skip_first_arg_len_eq_1_step: assumes "inv_tm_skip_first_arg_len_eq_1 n cf" shows "inv_tm_skip_first_arg_len_eq_1 n (step0 cf tm_skip_first_arg)" proof (cases cf) case (fields s l r) then have cf_cases: "cf = (s, l, r)" . show "inv_tm_skip_first_arg_len_eq_1 n (step0 cf tm_skip_first_arg)" proof (rule tm_skip_first_arg_len_eq_1_cases) from cf_cases and assms show "inv_tm_skip_first_arg_len_eq_1 n (s, l, r)" by auto next assume "s = 0" with cf_cases and assms show ?thesis by (auto simp add: tm_skip_first_arg_def) next assume "s = 1" show ?thesis proof - have "inv_tm_skip_first_arg_len_eq_1 n (step0 (1, l, r) tm_skip_first_arg)" proof (cases r) case Nil then have "r = []" . with assms and cf_cases and \s = 1\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) next case (Cons a rs) then have "r = a # rs" . show ?thesis proof (cases a) next case Bk then have "a = Bk" . with assms and \r = a # rs\ and cf_cases and \s = 1\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) next case Oc then have "a = Oc" . with assms and \r = a # rs\ and cf_cases and \s = 1\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) qed qed with cf_cases and \s=1\ show ?thesis by auto qed next assume "s = 2" show ?thesis proof - have "inv_tm_skip_first_arg_len_eq_1 n (step0 (2, l, r) tm_skip_first_arg)" proof (cases r) case Nil then have "r = []" . with assms and cf_cases and \s = 2\ have "inv_tm_skip_first_arg_len_eq_1_s2 n (l, r)" by auto then have "(\n1 n2. l = Oc \ (Suc n1) \ r = Oc \ n2 \ Suc n1 + n2 = Suc n)" by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) then obtain n1 n2 where w_n1_n2: "l = Oc \ (Suc n1) \ r = Oc \ n2 \ Suc n1 + n2 = Suc n" by blast with \r = []\ have "n2 = 0" by auto then have "step0 (2, Oc \ (Suc n1), Oc \ n2) tm_skip_first_arg = (3, Bk # Oc \ (Suc n1), [])" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover with \n2 = 0\ and w_n1_n2 have "inv_tm_skip_first_arg_len_eq_1 n (3, Bk # Oc \ (Suc n1), [])" by fastforce ultimately show ?thesis using w_n1_n2 by auto next case (Cons a rs) then have "r = a # rs" . show ?thesis proof (cases a) case Bk (* not possible due to invariant in state 2 *) then have "a = Bk" . with assms and \r = a # rs\ and cf_cases and \s = 2\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) next case Oc then have "a = Oc" . with assms and cf_cases and \s = 2\ have "inv_tm_skip_first_arg_len_eq_1_s2 n (l, r)" by auto then have "\n1 n2. l = Oc \ (Suc n1) \ r = Oc \ n2 \ Suc n1 + n2 = Suc n" by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) then obtain n1 n2 where w_n1_n2: "l = Oc \ (Suc n1) \ r = Oc \ n2 \ Suc n1 + n2 = Suc n" by blast with \r = a # rs\ and \a = Oc\ have "Oc # rs = Oc \ n2" by auto then have "n2 > 0" by (meson Cons_replicate_eq) then have "step0 (2, Oc \ (Suc n1), Oc \ n2) tm_skip_first_arg = (2, Oc # Oc \ (Suc n1), Oc \ (n2-1))" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_eq_1 n (2, Oc # Oc \ (Suc n1), Oc \ (n2-1))" proof - from \n2 > 0\ and w_n1_n2 have "Oc # Oc \ (Suc n1) = Oc \ (Suc (Suc n1)) \ Oc \ (n2-1) = Oc \ (n2-1) \ Suc (Suc n1) + (n2-1) = Suc n" by auto then have "(\n1' n2'. Oc # Oc \ (Suc n1) = Oc \ (Suc n1') \ Oc \ (n2-1) = Oc \ n2' \ Suc n1' + n2' = Suc n)" by auto then show "inv_tm_skip_first_arg_len_eq_1 n (2, Oc # Oc \ (Suc n1), Oc \ (n2-1))" by auto qed ultimately show ?thesis using assms and \r = a # rs\ and cf_cases and \s = 2\ and w_n1_n2 by auto qed qed with cf_cases and \s=2\ show ?thesis by auto qed next assume "s = 3" show ?thesis proof - have "inv_tm_skip_first_arg_len_eq_1 n (step0 (3, l, r) tm_skip_first_arg)" proof (cases r) case Nil then have "r = []" . with assms and cf_cases and \s = 3\ have "inv_tm_skip_first_arg_len_eq_1_s3 n (l, r)" by auto then have "l = Bk # Oc \ (Suc n) \ r = []" by auto then have "step0 (3, Bk # Oc \ (Suc n), []) tm_skip_first_arg = (4, Oc \ (Suc n), [Bk])" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_eq_1 n (4, Oc \ (Suc n), [Bk])" by fastforce ultimately show ?thesis using \l = Bk # Oc \ (Suc n) \ r = []\ by auto next case (Cons a rs) (* not possible due to invariant in state 3 *) then have "r = a # rs" . with assms and cf_cases and \s = 3\ have "inv_tm_skip_first_arg_len_eq_1_s3 n (l, r)" by auto then have "l = Bk # Oc \ (Suc n) \ r = []" by auto with \r = a # rs\ have False by auto then show ?thesis by auto qed with cf_cases and \s=3\ show ?thesis by auto qed next assume "s = 4" show ?thesis proof - have "inv_tm_skip_first_arg_len_eq_1 n (step0 (4, l, r) tm_skip_first_arg)" proof (cases r) case Nil (* not possible due to invariant in state 4 *) then have "r = []" . with assms and cf_cases and \s = 4\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) next case (Cons a rs) then have "r = a # rs" . show ?thesis proof (cases a) next case Bk then have "a = Bk" . with assms and \r = a # rs\ and cf_cases and \s = 4\ have "inv_tm_skip_first_arg_len_eq_1_s4 n (l, r)" by auto then have "l = Oc \ (Suc n) \ r = [Bk]" by auto then have F0: "step0 (4, Oc \ (Suc n), [Bk]) tm_skip_first_arg = (5, Oc \ n, Oc \ (Suc 0) @ [Bk])" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_eq_1_s5 n (Oc \ n, Oc \ (Suc 0) @ [Bk])" proof (cases n) case 0 then have "n=0" . then have "inv_tm_skip_first_arg_len_eq_1_s5 0 ([], Oc \ (Suc 0) @ [Bk])" by auto moreover with \n=0\ have "(5, Oc \ n, Oc \ (Suc 0) @ [Bk]) = (5, [], Oc \ (Suc 0) @ [Bk])" by auto ultimately show ?thesis by auto next case (Suc n') then have "n = Suc n'" . then have "(5, Oc \ n, Oc \ (Suc 0) @ [Bk]) = (5, Oc \ Suc n', Oc \ (Suc 0) @ [Bk])" by auto with \n=Suc n'\ have "Suc n' + Suc 0 = Suc n" by arith then have "(Oc \ Suc n' = Oc \ Suc n' \ Oc \ (Suc 0) @ [Bk] = Oc \ Suc 0 @ [Bk] \ Suc n' + Suc 0 = Suc n)" by auto with \(5, Oc \ n, Oc \ (Suc 0) @ [Bk]) = (5, Oc \ Suc n', Oc \ (Suc 0) @ [Bk])\ show ?thesis by (simp add: Suc \Suc n' + Suc 0 = Suc n\) qed then have "inv_tm_skip_first_arg_len_eq_1 n (5, Oc \ n, Oc \ (Suc 0) @ [Bk])" by auto ultimately show ?thesis using \l = Oc \ (Suc n) \ r = [Bk]\ by auto next case Oc (* not possible due to invariant in state 4 *) then have "a = Oc" . with assms and \r = a # rs\ and cf_cases and \s = 4\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) qed qed with cf_cases and \s=4\ show ?thesis by auto qed next assume "s = 5" show ?thesis proof - have "inv_tm_skip_first_arg_len_eq_1 n (step0 (5, l, r) tm_skip_first_arg)" proof (cases r) case Nil (* not possible due to invariant in state 5 *) then have "r = []" . with assms and cf_cases and \s = 5\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) next case (Cons a rs) then have "r = a # rs" . show ?thesis proof (cases a) case Bk then have "a = Bk" . with assms and \r = a # rs\ and cf_cases and \s = 5\ have "inv_tm_skip_first_arg_len_eq_1_s5 n (l, r)" by auto then have "\n1 n2. (l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) \ (l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n)" by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) then obtain n1 n2 where w_n1_n2: "(l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) \ (l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n)" by blast with \a = Bk\ and \r = a # rs\ have "l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n" by auto then have "step0 (5, [], Bk#Oc \ Suc n2 @ [Bk]) tm_skip_first_arg = (0, [Bk], Oc \ Suc n2 @ [Bk])" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_eq_1 n (0, [Bk], Oc \ Suc n @ [Bk])" proof - have "inv_tm_skip_first_arg_len_eq_1_s0 n ([Bk], Oc \ Suc n @ [Bk])" by (simp) then show "inv_tm_skip_first_arg_len_eq_1 n (0, [Bk], Oc \ Suc n @ [Bk])" by auto qed ultimately show ?thesis using assms and \a = Bk\ and \r = a # rs\ and cf_cases and \s = 5\ and \l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n\ by (simp) next case Oc then have "a = Oc" . with assms and \r = a # rs\ and cf_cases and \s = 5\ have "inv_tm_skip_first_arg_len_eq_1_s5 n (l, r)" by auto then have "\n1 n2. (l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) \ (l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n)" by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) then obtain n1 n2 where w_n1_n2: "(l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) \ (l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n)" by blast with \a = Oc\ and \r = a # rs\ have "(l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n)" by auto then show ?thesis proof assume "l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n" then have "step0 (5, l, r) tm_skip_first_arg = (5, Oc \ n1 , Oc \ Suc (Suc n2) @ [Bk])" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_eq_1 n (5, Oc \ n1 , Oc \ Suc (Suc n2) @ [Bk])" proof - from \l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n\ have "inv_tm_skip_first_arg_len_eq_1_s5 n (Oc \ n1, Oc \ Suc (Suc n2) @ [Bk])" by (cases n1) auto then show "inv_tm_skip_first_arg_len_eq_1 n (5, Oc \ n1 , Oc \ Suc (Suc n2) @ [Bk])" by auto qed ultimately show "inv_tm_skip_first_arg_len_eq_1 n (step0 (5, l, r) tm_skip_first_arg)" by auto next assume "l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n" then have "step0 (5, l, r) tm_skip_first_arg = (5, [], Bk # Oc \ Suc n2 @ [Bk])" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_eq_1 n (step0 (5, l, r) tm_skip_first_arg)" proof - from \l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n\ have "inv_tm_skip_first_arg_len_eq_1_s5 n (l, r)" by simp then show "inv_tm_skip_first_arg_len_eq_1 n (step0 (5, l, r) tm_skip_first_arg)" using \l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n\ and \step0 (5, l, r) tm_skip_first_arg = (5, [], Bk # Oc \ Suc n2 @ [Bk])\ by simp qed ultimately show ?thesis using assms and \a = Oc\ and \r = a # rs\ and cf_cases and \s = 5\ and \l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n\ by (simp ) qed qed qed with cf_cases and \s=5\ show ?thesis by auto qed qed qed lemma inv_tm_skip_first_arg_len_eq_1_steps: assumes "inv_tm_skip_first_arg_len_eq_1 n cf" shows "inv_tm_skip_first_arg_len_eq_1 n (steps0 cf tm_skip_first_arg stp)" proof (induct stp) case 0 with assms show ?case by (auto simp add: inv_tm_skip_first_arg_len_eq_1_step step.simps steps.simps) next case (Suc stp) with assms show ?case using inv_tm_skip_first_arg_len_eq_1_step step_red by auto qed lemma tm_skip_first_arg_len_eq_1_partial_correctness: assumes "\stp. is_final (steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp)" shows "\ \tap. tap = ([], <[n::nat]>) \ tm_skip_first_arg \ \tap. tap = ([Bk], <[n::nat]> @[Bk]) \" proof (rule Hoare_consequence) show "(\tap. tap = ([], <[n::nat]>)) \ (\tap. tap = ([], <[n::nat]>))" by auto next show "inv_tm_skip_first_arg_len_eq_1_s0 n \ (\tap. tap = ([Bk], <[n::nat]> @ [Bk]))" by (simp add: assert_imp_def tape_of_list_def tape_of_nat_def) next show "\\tap. tap = ([], <[n]>)\ tm_skip_first_arg \inv_tm_skip_first_arg_len_eq_1_s0 n\" proof (rule Hoare_haltI) fix l::"cell list" fix r:: "cell list" assume major: "(l, r) = ([], <[n]>)" show "\stp. is_final (steps0 (1, l, r) tm_skip_first_arg stp) \ inv_tm_skip_first_arg_len_eq_1_s0 n holds_for steps0 (1, l, r) tm_skip_first_arg stp" proof - from major and assms have "\stp. is_final (steps0 (1, l, r) tm_skip_first_arg stp)" by auto then obtain stp where w_stp: "is_final (steps0 (1, l, r) tm_skip_first_arg stp)" by blast moreover have "inv_tm_skip_first_arg_len_eq_1_s0 n holds_for steps0 (1, l, r) tm_skip_first_arg stp" proof - have "inv_tm_skip_first_arg_len_eq_1 n (1, l, r)" by (simp add: major tape_of_list_def tape_of_nat_def) then have "inv_tm_skip_first_arg_len_eq_1 n (steps0 (1, l, r) tm_skip_first_arg stp)" using inv_tm_skip_first_arg_len_eq_1_steps by auto then show ?thesis - by (smt holds_for.elims(3) inv_tm_skip_first_arg_len_eq_1.simps is_final_eq w_stp) + by (smt (verit) holds_for.elims(3) inv_tm_skip_first_arg_len_eq_1.simps is_final_eq w_stp) qed ultimately show ?thesis by auto qed qed qed (* --- now, we prove termination of tm_skip_first_arg on singleton lists --- *) (* Lexicographic orders (See List.measures) quote: "These are useful for termination proofs" lemma in_measures[simp]: "(x, y) \ measures [] = False" "(x, y) \ measures (f # fs) = (f x < f y \ (f x = f y \ (x, y) \ measures fs))" *) (* Assemble a lexicographic measure function *) definition measure_tm_skip_first_arg_len_eq_1 :: "(config \ config) set" where "measure_tm_skip_first_arg_len_eq_1 = measures [ \(s, l, r). (if s = 0 then 0 else 5 - s), \(s, l, r). (if s = 2 then length r else 0), \(s, l, r). (if s = 5 then length l + (if hd r = Oc then 2 else 1) else 0) ]" lemma wf_measure_tm_skip_first_arg_len_eq_1: "wf measure_tm_skip_first_arg_len_eq_1" unfolding measure_tm_skip_first_arg_len_eq_1_def by (auto) lemma measure_tm_skip_first_arg_len_eq_1_induct [case_names Step]: "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_tm_skip_first_arg_len_eq_1\ \ \n. P (f n)" using wf_measure_tm_skip_first_arg_len_eq_1 by (metis wf_iff_no_infinite_down_chain) (* Machine tm_skip_first_arg always halts on a singleton list *) lemma tm_skip_first_arg_len_eq_1_halts: "\stp. is_final (steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp)" proof (induct rule: measure_tm_skip_first_arg_len_eq_1_induct) case (Step stp) then have not_final: "\ is_final (steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp)" . have INV: "inv_tm_skip_first_arg_len_eq_1 n (steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp)" proof (rule_tac inv_tm_skip_first_arg_len_eq_1_steps) show "inv_tm_skip_first_arg_len_eq_1 n (1, [], <[n::nat]>)" by (simp add: tape_of_list_def tape_of_nat_def ) qed have SUC_STEP_RED: "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp) tm_skip_first_arg" by (rule step_red) show "( steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp), steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp ) \ measure_tm_skip_first_arg_len_eq_1" proof (cases "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp") case (fields s l r) then have cf_cases: "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (s, l, r)" . show ?thesis proof (rule tm_skip_first_arg_len_eq_1_cases) from INV and cf_cases show "inv_tm_skip_first_arg_len_eq_1 n (s, l, r)" by auto next assume "s=0" (* not possible *) with cf_cases not_final show ?thesis by auto next assume "s=1" show ?thesis proof (cases r) case Nil then have "r = []" . with cf_cases and \s=1\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (1, l, [])" by auto with INV have False by auto (* not possible due to invariant in s=1 *) then show ?thesis by auto next case (Cons a rs) then have "r = a # rs" . show ?thesis proof (cases "a") case Bk then have "a=Bk" . with cf_cases and \s=1\ and \r = a # rs\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (1, l, Bk#rs)" by auto with INV have False by auto (* not possible due to invariant in s=1 *) then show ?thesis by auto next case Oc then have "a=Oc" . with cf_cases and \s=1\ and \r = a # rs\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (1, l, Oc#rs)" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (1, l, Oc#rs) tm_skip_first_arg" by auto also have "... = (2,Oc#l,rs)" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (2,Oc#l,rs)" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (1, l, Oc#rs)\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) qed qed next assume "s=2" show ?thesis proof - from cf_cases and \s=2\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (2, l, r)" by auto with cf_cases and \s=2\ and INV have "(\n1 n2. l = Oc \ (Suc n1) \ r = Oc \ n2 \ Suc n1 + n2 = Suc n)" by auto then have "(\n2. r = Oc \ n2)" by blast then obtain n2 where w_n2: "r = Oc \ n2" by blast show ?thesis proof (cases n2) case 0 then have "n2 = 0" . with w_n2 have "r = []" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (2, l, r)\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (2, l, [])" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (2, l, []) tm_skip_first_arg" by auto also have "... = (3,Bk#l,[])" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (3,Bk#l,[])" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (2, l, [])\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) next case (Suc n2') then have "n2 = Suc n2'" . with w_n2 have "r = Oc \ Suc n2'" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (2, l, r)\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (2, l, Oc#Oc \ n2')" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (2, l, Oc#Oc \ n2') tm_skip_first_arg" by auto also have "... = (2,Oc#l,Oc \ n2')" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (2,Oc#l,Oc \ n2')" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (2, l, Oc#Oc \ n2')\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) qed qed next assume "s=3" show ?thesis proof - from cf_cases and \s=3\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (3, l, r)" by auto with cf_cases and \s=3\ and INV have "l = Bk # Oc \ (Suc n) \ r = []" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (3, l, r)\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (3, Bk # Oc \ (Suc n), [])" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (3, Bk # Oc \ (Suc n), []) tm_skip_first_arg" by auto also have "... = (4,Oc \ (Suc n),[Bk])" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (4,Oc \ (Suc n),[Bk])" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (3, Bk # Oc \ (Suc n), [])\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) qed next assume "s=4" show ?thesis proof - from cf_cases and \s=4\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (4, l, r)" by auto with cf_cases and \s=4\ and INV have "l = Oc \ (Suc n) \ r = [Bk]" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (4, l, r)\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (4, Oc \ (Suc n), [Bk])" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (4, Oc \ (Suc n), [Bk]) tm_skip_first_arg" by auto also have "... = (5,Oc \ n,[Oc, Bk])" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (5,Oc \ n,[Oc, Bk])" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (4, Oc \ (Suc n), [Bk])\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) qed next assume "s=5" show ?thesis proof - from cf_cases and \s=5\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, l, r)" by auto with cf_cases and \s=5\ and INV have "(\n1 n2. (l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) \ (l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) )" by auto then obtain n1 n2 where w_n1_n2: "(l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n) \ (l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n) \ (l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n)" by blast then show ?thesis proof assume "l = Oc \ Suc n1 \ r = Oc \ Suc n2 @ [Bk] \ Suc n1 + Suc n2 = Suc n" with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, l, r)\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, Oc \ Suc n1, Oc \ Suc n2 @ [Bk])" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (5, Oc \ Suc n1, Oc \ Suc n2 @ [Bk]) tm_skip_first_arg" by auto also have "... = (5,Oc \ n1,Oc#Oc \ Suc n2 @ [Bk])" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (5,Oc \ n1,Oc#Oc \ Suc n2 @ [Bk])" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, Oc \ Suc n1, Oc \ Suc n2 @ [Bk])\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) next assume "l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n \ l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n" then show ?thesis proof assume "l = [] \ r = Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n" with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, l, r)\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, [], Oc \ Suc n2 @ [Bk])" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (5, [], Oc \ Suc n2 @ [Bk]) tm_skip_first_arg" by auto also have "... = (5,[],Bk#Oc \ Suc n2 @ [Bk])" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (5,[],Bk#Oc \ Suc n2 @ [Bk])" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, [], Oc \ Suc n2 @ [Bk])\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) next assume "l = [] \ r = Bk # Oc \ Suc n2 @ [Bk] \ Suc n2 = Suc n" with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, l, r)\ have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, [], Bk # Oc \ Suc n2 @ [Bk])" by auto with SUC_STEP_RED have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = step0 (5, [], Bk # Oc \ Suc n2 @ [Bk]) tm_skip_first_arg" by auto also have "... = (0,[Bk], Oc \ Suc n2 @ [Bk])" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], <[n::nat]>) tm_skip_first_arg (Suc stp) = (0,[Bk], Oc \ Suc n2 @ [Bk])" by auto with \steps0 (1, [], <[n::nat]>) tm_skip_first_arg stp = (5, [], Bk # Oc \ Suc n2 @ [Bk])\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_eq_1_def) qed qed qed qed qed qed lemma tm_skip_first_arg_len_eq_1_total_correctness: "\ \tap. tap = ([], <[n::nat]>)\ tm_skip_first_arg \ \tap. tap = ([Bk], <[n::nat]> @[Bk])\" proof (rule tm_skip_first_arg_len_eq_1_partial_correctness) show "\stp. is_final (steps0 (1, [], <[n]>) tm_skip_first_arg stp)" using tm_skip_first_arg_len_eq_1_halts by auto qed lemma tm_skip_first_arg_len_eq_1_total_correctness': "length nl = 1 \ \\tap. tap = ([], )\ tm_skip_first_arg \ \tap. tap = ([Bk], <[hd nl]> @[Bk])\" proof - assume "length nl = 1" then have "nl = [hd nl]" by (metis One_nat_def diff_Suc_1 length_0_conv length_greater_0_conv length_tl list.distinct(1) list.expand list.sel(1) list.sel(3) list.size(3) zero_neq_one) moreover have "\ \tap. tap = ([], <[hd nl]>)\ tm_skip_first_arg \ \tap. tap = ([Bk], <[hd nl]> @[Bk])\" by (rule tm_skip_first_arg_len_eq_1_total_correctness) ultimately show ?thesis by (simp add: Hoare_haltE Hoare_haltI tape_of_list_def) qed (* ERROR cases: length nl > 1 *) (* ctxrunTM tm_skip_first_arg (1,[],[Oc, Bk, Oc]) 0: [] _1_ [Oc,Bk,Oc] => 1: [Oc] _2_ [Bk,Oc] => 2: [Oc,Bk] _3_ [Oc] => 3: [Oc,Bk] _0_ [Oc] ctxrunTM tm_skip_first_arg (1,[],[Oc,Oc, Bk, Oc,Oc]) 0: [] _1_ [Oc,Oc,Bk,Oc,Oc] => 1: [Oc] _2_ [Oc,Bk,Oc,Oc] => 2: [Oc,Oc] _2_ [Bk,Oc,Oc] => 3: [Oc,Oc,Bk] _3_ [Oc,Oc] => 4: [Oc,Oc,Bk] _0_ [Oc,Oc] ctxrunTM tm_skip_first_arg (1,[],[Oc,Oc,Oc, Bk, Oc,Oc]) 0: [] _1_ [Oc,Oc,Oc,Bk,Oc,Oc] => 1: [Oc] _2_ [Oc,Oc,Bk,Oc,Oc] => 2: [Oc,Oc] _2_ [Oc,Bk,Oc,Oc] => 3: [Oc,Oc,Oc] _2_ [Bk,Oc,Oc] => 4: [Oc,Oc,Oc,Bk] _3_ [Oc,Oc] => 5: [Oc,Oc,Oc,Bk] _0_ [Oc,Oc] ctxrunTM tm_skip_first_arg (1, [], [Oc,Oc,Bk, Oc,Bk, Oc,Oc]) 0: [] _1_ [Oc,Oc,Bk,Oc,Bk,Oc,Oc] => 1: [Oc] _2_ [Oc,Bk,Oc,Bk,Oc,Oc] => 2: [Oc,Oc] _2_ [Bk,Oc,Bk,Oc,Oc] => 3: [Bk,Oc,Oc] _3_ [Oc,Bk,Oc,Oc] => 4: [Bk,Oc,Oc] _0_ [Oc,Bk,Oc,Oc] *) fun inv_tm_skip_first_arg_len_gt_1_s0 :: "nat \ nat list \ tape \ bool" and inv_tm_skip_first_arg_len_gt_1_s1 :: "nat \ nat list\ tape \ bool" and inv_tm_skip_first_arg_len_gt_1_s2 :: "nat \ nat list\ tape \ bool" and inv_tm_skip_first_arg_len_gt_1_s3 :: "nat \ nat list\ tape \ bool" where "inv_tm_skip_first_arg_len_gt_1_s1 n ns (l, r) = ( l = [] \ r = Oc \ Suc n @ [Bk] @ () )" | "inv_tm_skip_first_arg_len_gt_1_s2 n ns (l, r) = (\n1 n2. l = Oc \ (Suc n1) \ r = Oc \ n2 @ [Bk] @ () \ Suc n1 + n2 = Suc n)" | "inv_tm_skip_first_arg_len_gt_1_s3 n ns (l, r) = ( l = Bk # Oc \ (Suc n) \ r = () )" | "inv_tm_skip_first_arg_len_gt_1_s0 n ns (l, r) = ( l = Bk# Oc \ (Suc n) \ r = () )" fun inv_tm_skip_first_arg_len_gt_1 :: "nat \ nat list \ config \ bool" where "inv_tm_skip_first_arg_len_gt_1 n ns (s, tap) = (if s = 0 then inv_tm_skip_first_arg_len_gt_1_s0 n ns tap else if s = 1 then inv_tm_skip_first_arg_len_gt_1_s1 n ns tap else if s = 2 then inv_tm_skip_first_arg_len_gt_1_s2 n ns tap else if s = 3 then inv_tm_skip_first_arg_len_gt_1_s3 n ns tap else False)" lemma tm_skip_first_arg_len_gt_1_cases: fixes s::nat assumes "inv_tm_skip_first_arg_len_gt_1 n ns (s,l,r)" and "s=0 \ P" and "s=1 \ P" and "s=2 \ P" and "s=3 \ P" and "s=4 \ P" and "s=5 \ P" shows "P" proof - have "s < 6" proof (rule ccontr) assume "\ s < 6" with \inv_tm_skip_first_arg_len_gt_1 n ns (s,l,r)\ show False by auto qed then have "s = 0 \ s = 1 \ s = 2 \ s = 3 \ s = 4 \ s = 5" by auto with assms show ?thesis by auto qed lemma inv_tm_skip_first_arg_len_gt_1_step: assumes "length ns > 0" and "inv_tm_skip_first_arg_len_gt_1 n ns cf" shows "inv_tm_skip_first_arg_len_gt_1 n ns (step0 cf tm_skip_first_arg)" proof (cases cf) case (fields s l r) then have cf_cases: "cf = (s, l, r)" . show "inv_tm_skip_first_arg_len_gt_1 n ns (step0 cf tm_skip_first_arg)" proof (rule tm_skip_first_arg_len_gt_1_cases) from cf_cases and assms show "inv_tm_skip_first_arg_len_gt_1 n ns (s, l, r)" by auto next assume "s = 0" with cf_cases and assms show ?thesis by (auto simp add: tm_skip_first_arg_def) next assume "s = 4" with cf_cases and assms show ?thesis by (auto simp add: tm_skip_first_arg_def) next assume "s = 5" with cf_cases and assms show ?thesis by (auto simp add: tm_skip_first_arg_def) next assume "s = 1" with cf_cases and assms have "l = [] \ r = Oc \ Suc n @ [Bk] @ ()" by auto with assms and cf_cases and \s = 1\ show ?thesis by (auto simp add: tm_skip_first_arg_def step.simps steps.simps) next assume "s = 2" with cf_cases and assms have "(\n1 n2. l = Oc \ (Suc n1) \ r = Oc \ n2 @ [Bk] @ () \ Suc n1 + n2 = Suc n)" by auto then obtain n1 n2 where w_n1_n2: "l = Oc \ (Suc n1) \ r = Oc \ n2 @ [Bk] @ () \ Suc n1 + n2 = Suc n" by blast show ?thesis proof (cases n2) case 0 then have "n2 = 0" . with w_n1_n2 have "l = Oc \ (Suc n1) \ r = [Bk] @ () \ Suc n1 = Suc n" by auto then have "step0 (2, Oc \ (Suc n1), [Bk] @ ()) tm_skip_first_arg = (3, Bk # Oc \ (Suc n1), ())" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_gt_1 n ns (3, Bk # Oc \ (Suc n1), ())" proof - from \l = Oc \ (Suc n1) \ r = [Bk] @ () \ Suc n1 = Suc n\ have "Oc \ (Suc n1) = Oc \ (Suc n1) \ Bk # Oc \ (Suc n1) = Bk#Oc#Oc \ n1 \ Suc n1 + 0 = Suc n" by auto then show "inv_tm_skip_first_arg_len_gt_1 n ns (3, Bk # Oc \ (Suc n1), ())" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 2\ and w_n1_n2 by auto next case (Suc n2') then have "n2 = Suc n2'" . with w_n1_n2 have "l = Oc \ (Suc n1) \ r = Oc \ Suc n2' @ [Bk] @ () \ Suc n1 + n2 = Suc n" by auto then have "step0 (2, Oc \ (Suc n1), Oc \ Suc n2' @ [Bk] @ ()) tm_skip_first_arg = (2, Oc # Oc \ (Suc n1), Oc \ n2' @ [Bk] @ ())" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_skip_first_arg_len_gt_1 n ns (2, Oc # Oc \ (Suc n1), Oc \ n2' @ [Bk] @ ())" proof - from \l = Oc \ (Suc n1) \ r = Oc \ Suc n2' @ [Bk] @ () \ Suc n1 + n2 = Suc n\ have "Oc # Oc \ (Suc n1) = Oc \ Suc (Suc n1) \ Oc \ n2' @ [Bk] @ () = Oc \ n2' @ [Bk] @ () \ Suc (Suc n1) + n2' = Suc n" by (simp add: Suc add_Suc_shift) then show "inv_tm_skip_first_arg_len_gt_1 n ns (2, Oc # Oc \ (Suc n1), Oc \ n2' @ [Bk] @ ())" by force qed ultimately show ?thesis using assms and cf_cases and \s = 2\ and w_n1_n2 using \l = Oc \ (Suc n1) \ r = Oc \ Suc n2' @ [Bk] @ () \ Suc n1 + n2 = Suc n\ by force qed next assume "s = 3" with cf_cases and assms have unpackedINV: "l = Bk # Oc \ (Suc n) \ r = ()" by auto moreover with \length ns > 0\ have "(ns::nat list) \ [] \ hd () = Oc" using numeral_list_head_is_Oc by force moreover from this have " = Oc#(tl ())" by (metis append_Nil list.exhaust_sel tape_of_list_empty unique_Bk_postfix_numeral_list_Nil) ultimately have "step0 (3, Bk # Oc \ (Suc n), ()) tm_skip_first_arg = (0, Bk # Oc \ (Suc n), ())" proof - from \ = Oc # tl ()\ have "step0 (3, Bk # Oc \ (Suc n), ()) tm_skip_first_arg = step0 (3, Bk # Oc \ (Suc n), Oc # tl ()) tm_skip_first_arg" by auto also have "... = (0, Bk # Oc \ (Suc n), Oc # tl ())" by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12) also with \ = Oc # tl ()\ have "... = (0, Bk # Oc \ (Suc n), ())" by auto finally show "step0 (3, Bk # Oc \ (Suc n), ()) tm_skip_first_arg = (0, Bk # Oc \ (Suc n), ())" by auto qed moreover with \l = Bk # Oc \ (Suc n) \ r = ()\ have "inv_tm_skip_first_arg_len_gt_1 n ns (0, Bk # Oc \ (Suc n), ())" by auto ultimately show ?thesis using assms and cf_cases and \s = 3\ and unpackedINV by auto qed qed lemma inv_tm_skip_first_arg_len_gt_1_steps: assumes "length ns > 0" and "inv_tm_skip_first_arg_len_gt_1 n ns cf" shows "inv_tm_skip_first_arg_len_gt_1 n ns (steps0 cf tm_skip_first_arg stp)" proof (induct stp) case 0 with assms show ?case by (auto simp add: inv_tm_skip_first_arg_len_gt_1_step step.simps steps.simps) next case (Suc stp) with assms show ?case using inv_tm_skip_first_arg_len_gt_1_step step_red by auto qed lemma tm_skip_first_arg_len_gt_1_partial_correctness: assumes "\stp. is_final (steps0 (1, [], Oc \ Suc n @ [Bk] @ () ) tm_skip_first_arg stp)" and "0 < length ns" shows "\\tap. tap = ([], Oc \ Suc n @ [Bk] @ () ) \ tm_skip_first_arg \ \tap. tap = (Bk# Oc \ Suc n, () ) \" proof (rule Hoare_consequence) show " (\tap. tap = ([], Oc \ Suc n @ [Bk] @ ())) \ (\tap. tap = ([], Oc \ Suc n @ [Bk] @ ()))" by (simp add: assert_imp_def tape_of_nat_def) next show "inv_tm_skip_first_arg_len_gt_1_s0 n ns \ (\tap. tap = (Bk # Oc \ Suc n, ))" using assert_imp_def inv_tm_skip_first_arg_len_gt_1_s0.simps rev_numeral tape_of_nat_def by auto next show "\\tap. tap = ([], Oc \ Suc n @ [Bk] @ )\ tm_skip_first_arg \inv_tm_skip_first_arg_len_gt_1_s0 n ns\" proof (rule Hoare_haltI) fix l::"cell list" fix r:: "cell list" assume major: "(l, r) = ([], Oc \ Suc n @ [Bk] @ )" show "\stp. is_final (steps0 (1, l, r) tm_skip_first_arg stp) \ inv_tm_skip_first_arg_len_gt_1_s0 n ns holds_for steps0 (1, l, r) tm_skip_first_arg stp" proof - from major and assms have "\stp. is_final (steps0 (1, l, r) tm_skip_first_arg stp)" by auto then obtain stp where w_stp: "is_final (steps0 (1, l, r) tm_skip_first_arg stp)" by blast moreover have "inv_tm_skip_first_arg_len_gt_1_s0 n ns holds_for steps0 (1, l, r) tm_skip_first_arg stp" proof - have "inv_tm_skip_first_arg_len_gt_1 n ns (1, l, r)" by (simp add: major tape_of_list_def tape_of_nat_def) with assms have "inv_tm_skip_first_arg_len_gt_1 n ns (steps0 (1, l, r) tm_skip_first_arg stp)" using inv_tm_skip_first_arg_len_gt_1_steps by auto then show ?thesis - by (smt holds_for.elims(3) inv_tm_skip_first_arg_len_gt_1.simps is_final_eq w_stp) + by (smt (verit) holds_for.elims(3) inv_tm_skip_first_arg_len_gt_1.simps is_final_eq w_stp) qed ultimately show ?thesis by auto qed qed qed (* --- now, we prove termination of tm_skip_first_arg on arguments containing more than one numeral --- *) (* Lexicographic orders (See List.measures) quote: "These are useful for termination proofs" lemma in_measures[simp]: "(x, y) \ measures [] = False" "(x, y) \ measures (f # fs) = (f x < f y \ (f x = f y \ (x, y) \ measures fs))" *) (* Assemble a lexicographic measure function *) definition measure_tm_skip_first_arg_len_gt_1 :: "(config \ config) set" where "measure_tm_skip_first_arg_len_gt_1 = measures [ \(s, l, r). (if s = 0 then 0 else 4 - s), \(s, l, r). (if s = 2 then length r else 0) ]" lemma wf_measure_tm_skip_first_arg_len_gt_1: "wf measure_tm_skip_first_arg_len_gt_1" unfolding measure_tm_skip_first_arg_len_gt_1_def by (auto) lemma measure_tm_skip_first_arg_len_gt_1_induct [case_names Step]: "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_tm_skip_first_arg_len_gt_1\ \ \n. P (f n)" using wf_measure_tm_skip_first_arg_len_gt_1 by (metis wf_iff_no_infinite_down_chain) lemma tm_skip_first_arg_len_gt_1_halts: "0 < length ns \ \stp. is_final (steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp)" proof - assume A: "0 < length ns" show "\stp. is_final (steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp)" proof (induct rule: measure_tm_skip_first_arg_len_gt_1_induct) case (Step stp) then have not_final: "\ is_final (steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp)" . have INV: "inv_tm_skip_first_arg_len_gt_1 n ns (steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp)" proof (rule_tac inv_tm_skip_first_arg_len_gt_1_steps) from A show "0 < length ns " . then show "inv_tm_skip_first_arg_len_gt_1 n ns (1, [], Oc \ Suc n @ [Bk] @ )" by (simp add: tape_of_list_def tape_of_nat_def) qed have SUC_STEP_RED: "steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg (Suc stp) = step0 (steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp) tm_skip_first_arg" by (rule step_red) show "( steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg (Suc stp), steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp ) \ measure_tm_skip_first_arg_len_gt_1" proof (cases "steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp") case (fields s l r2) then have cf_cases: "steps0 (1, [], Oc \ Suc n @ [Bk] @ ) tm_skip_first_arg stp = (s, l, r2)" . show ?thesis proof (rule tm_skip_first_arg_len_gt_1_cases) from INV and cf_cases show "inv_tm_skip_first_arg_len_gt_1 n ns (s, l, r2)" by auto next assume "s=0" (* not possible *) with cf_cases not_final show ?thesis by auto next assume "s=4" (* not possible *) with cf_cases not_final INV show ?thesis by auto next assume "s=5" (* not possible *) with cf_cases not_final INV show ?thesis by auto next assume "s=1" show ?thesis proof - from cf_cases and \s=1\ have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (1, l, r2)" by auto with cf_cases and \s=1\ and INV have unpackedINV: "l = [] \ r2 = Oc \ Suc n @ [Bk] @ ()" by auto with cf_cases and \s=1\ and SUC_STEP_RED have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = step0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg" by auto also have "... = (2,[Oc],Oc \ n @ [Bk] @ ())" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = (2,[Oc],Oc \ n @ [Bk] @ ())" by auto with \steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (1, l, r2)\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_gt_1_def) qed next assume "s=2" show ?thesis proof - from cf_cases and \s=2\ have "steps0 (1, [],Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (2, l, r2)" by auto with cf_cases and \s=2\ and INV have "(\n1 n2. l = Oc \ (Suc n1) \ r2 = Oc \ n2 @ [Bk] @ () \ Suc n1 + n2 = Suc n)" by auto then obtain n1 n2 where w_n1_n2: "l = Oc \ (Suc n1) \ r2 = Oc \ n2 @ [Bk] @ () \ Suc n1 + n2 = Suc n" by blast show ?thesis proof (cases n2) case 0 then have "n2 = 0" . with w_n1_n2 have "r2 = [Bk] @ ()" by auto with \steps0 (1, [],Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (2, l, r2)\ have "steps0 (1, [],Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (2, l, [Bk] @ ())" by auto with SUC_STEP_RED have "steps0 (1, [],Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = step0 (2, l, [Bk] @ ()) tm_skip_first_arg" by auto also have "... = (3,Bk#l,())" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [],Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = (3,Bk#l,())" by auto with \steps0 (1, [],Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (2, l, [Bk] @ ())\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_gt_1_def) next case (Suc n2') then have "n2 = Suc n2'" . with w_n1_n2 have "r2 = Oc \ Suc n2'@Bk#()" by auto with \steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (2, l, r2)\ have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (2, l, Oc \ Suc n2'@Bk#())" by auto with SUC_STEP_RED have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = step0 (2, l, Oc \ Suc n2'@Bk#()) tm_skip_first_arg" by auto also have "... = (2,Oc#l,Oc \ n2'@Bk#())" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = (2,Oc#l,Oc \ n2'@Bk#())" by auto with \steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (2, l, Oc \ Suc n2'@Bk#())\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_gt_1_def) qed qed next assume "s=3" show ?thesis proof - from cf_cases and \s=3\ have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (3, l, r2)" by auto with cf_cases and \s=3\ and INV have unpacked_INV: "l = Bk # Oc \ (Suc n) \ r2 = ()" by auto with \steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (3, l, r2)\ have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (3, Bk # Oc \ (Suc n), ())" by auto with SUC_STEP_RED have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = step0 (3, Bk # Oc \ (Suc n), ()) tm_skip_first_arg" by auto also have "... = (0, Bk # Oc \ (Suc n),())" proof - from \length ns > 0\ have "(ns::nat list) \ [] \ hd () = Oc" using numeral_list_head_is_Oc by force then have " = Oc#(tl ())" by (metis append_Nil list.exhaust_sel tape_of_list_empty unique_Bk_postfix_numeral_list_Nil) then have "step0 (3, Bk # Oc \ (Suc n), ()) tm_skip_first_arg = step0 (3, Bk # Oc \ (Suc n), Oc#(tl ())) tm_skip_first_arg" by auto also have "... = (0, Bk # Oc \ (Suc n), Oc#(tl ()))" by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps) also with \ = Oc#(tl ())\ have "... = (0, Bk # Oc \ (Suc n), )" by auto finally show "step0 (3, Bk # Oc \ (Suc n), ()) tm_skip_first_arg = (0, Bk # Oc \ (Suc n), )" by auto qed finally have "steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg (Suc stp) = (0, Bk # Oc \ (Suc n),())" by auto with \steps0 (1, [], Oc \ Suc n @ [Bk] @ ()) tm_skip_first_arg stp = (3, Bk # Oc \ (Suc n), ())\ show ?thesis by (auto simp add: measure_tm_skip_first_arg_len_gt_1_def) qed qed qed qed qed lemma tm_skip_first_arg_len_gt_1_total_correctness_pre: assumes "0 < length ns" shows "\\tap. tap = ([], Oc \ Suc n @ [Bk] @ () ) \ tm_skip_first_arg \ \tap. tap = (Bk# Oc \ Suc n, () ) \" proof (rule tm_skip_first_arg_len_gt_1_partial_correctness) from assms show "0 < length ns" . from assms show "\stp. is_final (steps0 (1, [], Oc \ Suc n @ [Bk] @ () ) tm_skip_first_arg stp)" using tm_skip_first_arg_len_gt_1_halts by auto qed lemma tm_skip_first_arg_len_gt_1_total_correctness: assumes "1 < length (nl::nat list)" shows "\\tap. tap = ([], )\ tm_skip_first_arg \ \tap. tap = (Bk# , ) \" proof - from assms have major: "(nl::nat list) = hd nl # tl nl" by (metis list.exhaust_sel list.size(3) not_one_less_zero) from assms have "tl nl \ []" using list_length_tl_neq_Nil by auto from assms have "(nl::nat list) \ []" by auto from \(nl::nat list) = hd nl # tl nl\ have " = " by auto also with \tl nl \ []\ have "... = @ [Bk] @ ()" by (simp add: tape_of_nat_list_cons_eq) also with \(nl::nat list) \ []\ have "... = Oc \ Suc (hd nl) @ [Bk] @ ()" using tape_of_nat_def by blast finally have " = Oc \ Suc (hd nl) @ [Bk] @ ()" by auto from \tl nl \ []\ have "0 < length (tl nl)" using length_greater_0_conv by blast with assms have "\\tap. tap = ([], Oc \ Suc (hd nl) @ [Bk] @ () ) \ tm_skip_first_arg \ \tap. tap = (Bk# Oc \ Suc (hd nl), () ) \" using tm_skip_first_arg_len_gt_1_total_correctness_pre by force with \ = Oc \ Suc (hd nl) @ [Bk] @ ()\ have "\\tap. tap = ([], ) \ tm_skip_first_arg \ \tap. tap = (Bk# Oc \ Suc (hd nl), () ) \" by force moreover have " = Oc \ Suc (hd nl)" by (simp add: tape_of_list_def tape_of_nat_def) ultimately show ?thesis by (simp add: rev_numeral rev_numeral_list tape_of_list_def ) qed (* ---------- tm_erase_right_then_dblBk_left ---------- *) (* We will prove: -- DO_NOTHING path, the eraser should do nothing lemma tm_erase_right_then_dblBk_left_dnp_total_correctness: "\ \tap. tap = ([], r ) \ tm_erase_right_then_dblBk_left \ \tap. tap = ([Bk,Bk], r ) \" -- ERASE path lemma tm_erase_right_then_dblBk_left_erp_total_correctness_one_arg: assumes "1 \ length (nl::nat list)" shows "\ \tap. tap = (Bk# rev(), ) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ () @ [Bk] @ Bk \ rex ) \" *) (* Test runs for the formulation of invariants: * * See file EngineeringOf_StrongCopy2.hs for more test runs * ctxrunTM tm_erase_right_then_dblBk_left (1, [Bk,Oc], [Oc]) 0: [Bk,Oc] _1_ [Oc] => 1: [Oc] _2_ [Bk,Oc] => 2: [] _3_ [Oc,Bk,Oc] => 3: [Oc] _5_ [Bk,Oc] => 4: [Bk,Oc] _6_ [Oc] => 5: [Bk,Oc] _6_ [Bk] => 6: [Bk,Bk,Oc] _7_ [] => 7: [Bk,Bk,Bk,Oc] _9_ [] => 8: [Bk,Bk,Oc] _10_ [Bk] => 9: [Bk,Oc] _10_ [Bk,Bk] => 10: [Oc] _10_ [Bk,Bk,Bk] => 11: [] _10_ [Oc,Bk,Bk,Bk] => 12: [] _11_ [Bk,Oc,Bk,Bk,Bk] => 13: [] _12_ [Bk,Bk,Oc,Bk,Bk,Bk] => 14: [] _0_ [Bk,Bk,Oc,Bk,Bk,Bk] ctxrunTM tm_erase_right_then_dblBk_left (1, [Bk,Oc,Oc], [Oc,Oc,Oc]) 0: [Bk,Oc,Oc] _1_ [Oc,Oc,Oc] => 1: [Oc,Oc] _2_ [Bk,Oc,Oc,Oc] => 2: [Oc] _3_ [Oc,Bk,Oc,Oc,Oc] => 3: [Oc,Oc] _5_ [Bk,Oc,Oc,Oc] => 4: [Bk,Oc,Oc] _6_ [Oc,Oc,Oc] => 5: [Bk,Oc,Oc] _6_ [Bk,Oc,Oc] => 6: [Bk,Bk,Oc,Oc] _7_ [Oc,Oc] => 7: [Bk,Bk,Oc,Oc] _8_ [Bk,Oc] => 8: [Bk,Bk,Bk,Oc,Oc] _7_ [Oc] => 9: [Bk,Bk,Bk,Oc,Oc] _8_ [Bk] => 10: [Bk,Bk,Bk,Bk,Oc,Oc] _7_ [] => 11: [Bk,Bk,Bk,Bk,Bk,Oc,Oc] _9_ [] => 12: [Bk,Bk,Bk,Bk,Oc,Oc] _10_ [Bk] => 13: [Bk,Bk,Bk,Oc,Oc] _10_ [Bk,Bk] => 14: [Bk,Bk,Oc,Oc] _10_ [Bk,Bk,Bk] => 15: [Bk,Oc,Oc] _10_ [Bk,Bk,Bk,Bk] => 16: [Oc,Oc] _10_ [Bk,Bk,Bk,Bk,Bk] => 17: [Oc] _10_ [Oc,Bk,Bk,Bk,Bk,Bk] => 18: [] _11_ [Oc,Oc,Bk,Bk,Bk,Bk,Bk] => 19: [] _11_ [Bk,Oc,Oc,Bk,Bk,Bk,Bk,Bk] => 20: [] _12_ [Bk,Bk,Oc,Oc,Bk,Bk,Bk,Bk,Bk] => 21: [] _0_ [Bk,Bk,Oc,Oc,Bk,Bk,Bk,Bk,Bk] *) (* Commented definition of tm_erase_right_then_dblBk_left let tm_erase_right_then_dblBk_left = from_raw [ -- 'check the left tape': (L, 2),(L, 2), -- 1 one step left (L, 3),(R, 5), -- 2 on Bk do one more step left, on Oc right towards 'check the right tape' (R, 4),(R, 5), -- 3 on 2nd Bk right to towards termination, on Oc right towards 'check the right tape' (R, 0),(R, 0), -- 4 one step right and terminate (R, 6),(R, 6), -- 5 one more step right to erase path -- 'check the right tape': (R, 7),(WB,6), -- 6 on Bk goto loop 'erase all to the right', on Oc erase -- (R, 9),(WB,8), -- 7 loop 'erase all to the right': on Bk leave loop, on Oc erase (R, 7),(R, 7), -- 8 move right and continue loop 'erase all to the right' (L,10),(WB,8), -- 9 on 2nd Bk start loop 'move to left until Oc', on Oc continue 'erase all to the right' -- (L,10),(L,11), -- 10 loop 'move to left until Oc': on Bk continue loop 'move to left until Oc', on Oc goto loop 'move left until DblBk' -- (L,12),(L,11), -- 11 loop 'move left until DblBk': on 1st Bk move once more to the left, on Oc continue loop 'move left until DblBk' (WB,0),(L,11) -- 12 on 2nd Bk stop, on Oc (possible in generic case) continue loop 'move left until DblBk' ] -- Note: -- In a more generic context, the 'move left until DblBk' may see an Oc in state 12. -- However, in our context of tm_check_for_one_arg there cannot be an Oc in state 12. *) definition tm_erase_right_then_dblBk_left :: "instr list" where "tm_erase_right_then_dblBk_left \ [ (L, 2),(L, 2), (L, 3),(R, 5), (R, 4),(R, 5), (R, 0),(R, 0), (R, 6),(R, 6), (R, 7),(WB,6), (R, 9),(WB,8), (R, 7),(R, 7), (L,10),(WB,8), (L,10),(L,11), (L,12),(L,11), (WB,0),(L,11) ]" (* Partial and total correctness for tm_erase_right_then_dblBk_left * * The leading 2 symbols on the left tape are used to differentiate! * * [Bk,Bk] \ DO NOTHING path * [Bk,Oc] \ ERASE path * *) (* ----------------------------------------------------------------- *) (* DO NOTHING path tm_erase_right_then_dblBk_left_dnp *) (* Sequence of states on the DO NOTHING path: 1,2,3 then 4, 0 *) (* ----------------------------------------------------------------- *) (* * "\ \tap. tap = ([], r ) \ * tm_erase_right_then_dblBk_left * \ \tap. tap = ([Bk,Bk], r ) \" *) (* Definition of invariants for the DO_NOTHING path *) fun inv_tm_erase_right_then_dblBk_left_dnp_s0 :: "(cell list) \ tape \ bool" and inv_tm_erase_right_then_dblBk_left_dnp_s1 :: "(cell list) \ tape \ bool" and inv_tm_erase_right_then_dblBk_left_dnp_s2 :: "(cell list) \ tape \ bool" and inv_tm_erase_right_then_dblBk_left_dnp_s3 :: "(cell list) \ tape \ bool" and inv_tm_erase_right_then_dblBk_left_dnp_s4 :: "(cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_dnp_s0 CR (l, r) = (l = [Bk, Bk] \ CR = r)" | "inv_tm_erase_right_then_dblBk_left_dnp_s1 CR (l, r) = (l = [] \ CR = r)" | "inv_tm_erase_right_then_dblBk_left_dnp_s2 CR (l, r) = (l = [] \ r = Bk#CR)" | "inv_tm_erase_right_then_dblBk_left_dnp_s3 CR (l, r) = (l = [] \ r = Bk#Bk#CR)" | "inv_tm_erase_right_then_dblBk_left_dnp_s4 CR (l, r) = (l = [Bk] \ r = Bk#CR)" fun inv_tm_erase_right_then_dblBk_left_dnp :: "(cell list) \ config \ bool" where "inv_tm_erase_right_then_dblBk_left_dnp CR (s, tap) = (if s = 0 then inv_tm_erase_right_then_dblBk_left_dnp_s0 CR tap else if s = 1 then inv_tm_erase_right_then_dblBk_left_dnp_s1 CR tap else if s = 2 then inv_tm_erase_right_then_dblBk_left_dnp_s2 CR tap else if s = 3 then inv_tm_erase_right_then_dblBk_left_dnp_s3 CR tap else if s = 4 then inv_tm_erase_right_then_dblBk_left_dnp_s4 CR tap else False)" lemma tm_erase_right_then_dblBk_left_dnp_cases: fixes s::nat assumes "inv_tm_erase_right_then_dblBk_left_dnp CR (s,l,r)" and "s=0 \ P" and "s=1 \ P" and "s=2 \ P" and "s=3 \ P" and "s=4 \ P" shows "P" proof - have "s < 5" proof (rule ccontr) assume "\ s < 5" with \inv_tm_erase_right_then_dblBk_left_dnp CR (s,l,r)\ show False by auto qed then have "s = 0 \ s = 1 \ s = 2 \ s = 3 \ s = 4" by auto with assms show ?thesis by auto qed lemma inv_tm_erase_right_then_dblBk_left_dnp_step: assumes "inv_tm_erase_right_then_dblBk_left_dnp CR cf" shows "inv_tm_erase_right_then_dblBk_left_dnp CR (step0 cf tm_erase_right_then_dblBk_left)" proof (cases cf) case (fields s l r) then have cf_cases: "cf = (s, l, r)" . show "inv_tm_erase_right_then_dblBk_left_dnp CR (step0 cf tm_erase_right_then_dblBk_left)" proof (rule tm_erase_right_then_dblBk_left_dnp_cases) from cf_cases and assms show "inv_tm_erase_right_then_dblBk_left_dnp CR (s, l, r)" by auto next assume "s = 0" with cf_cases and assms show ?thesis by (auto simp add: tm_erase_right_then_dblBk_left_def) next assume "s = 1" with cf_cases and assms have "l = []" by auto show ?thesis proof (cases r) case Nil then have "r = []" . with assms and cf_cases and \s = 1\ show ?thesis by (auto simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps) next case (Cons a rs) then have "r = a # rs" . show ?thesis proof (cases a) case Bk with assms and cf_cases and \s = 1\ and \r = a # rs\ show ?thesis by (auto simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps) next case Oc with assms and cf_cases and \s = 1\ and \r = a # rs\ show ?thesis by (auto simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps) qed qed next assume "s = 2" with cf_cases and assms have "l = [] \ r = Bk#CR" by auto then have "step0 (2, [], Bk#CR) tm_erase_right_then_dblBk_left = (3, [], Bk# Bk # CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_dnp CR (3, [], Bk# Bk # CR)" proof - from \l = [] \ r = Bk#CR\ show "inv_tm_erase_right_then_dblBk_left_dnp CR (3, [], Bk# Bk # CR)" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 2\ and \l = [] \ r = Bk#CR\ by auto next assume "s = 3" with cf_cases and assms have "l = [] \ r = Bk#Bk#CR" by auto then have "step0 (3, [], Bk#Bk#CR) tm_erase_right_then_dblBk_left = (4, [Bk], Bk # CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_dnp CR (4, [Bk], Bk # CR)" proof - from \l = [] \ r = Bk#Bk#CR\ show "inv_tm_erase_right_then_dblBk_left_dnp CR (4, [Bk], Bk # CR)" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 3\ and \l = [] \ r = Bk#Bk#CR\ by auto next assume "s = 4" with cf_cases and assms have "l = [Bk] \ r = Bk#CR" by auto then have "step0 (4, [Bk], Bk#CR) tm_erase_right_then_dblBk_left = (0, [Bk,Bk], CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_dnp CR (0, [Bk,Bk], CR)" proof - from \l = [Bk] \ r = Bk#CR\ show "inv_tm_erase_right_then_dblBk_left_dnp CR (0, [Bk,Bk], CR)" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 4\ and \l = [Bk] \ r = Bk#CR\ by auto qed qed lemma inv_tm_erase_right_then_dblBk_left_dnp_steps: assumes "inv_tm_erase_right_then_dblBk_left_dnp CR cf" shows "inv_tm_erase_right_then_dblBk_left_dnp CR (steps0 cf tm_erase_right_then_dblBk_left stp)" proof (induct stp) case 0 with assms show ?case by (auto simp add: inv_tm_erase_right_then_dblBk_left_dnp_step step.simps steps.simps) next case (Suc stp) with assms show ?case using inv_tm_erase_right_then_dblBk_left_dnp_step step_red by auto qed lemma tm_erase_right_then_dblBk_left_dnp_partial_correctness: assumes "\stp. is_final (steps0 (1, [], r) tm_erase_right_then_dblBk_left stp)" shows "\ \tap. tap = ([], r ) \ tm_erase_right_then_dblBk_left \ \tap. tap = ([Bk,Bk], r ) \" proof (rule Hoare_consequence) show "(\tap. tap = ([], r) ) \ (\tap. tap = ([], r) )" by auto next show "inv_tm_erase_right_then_dblBk_left_dnp_s0 r \ (\tap. tap = ([Bk,Bk], r ))" by (simp add: assert_imp_def tape_of_list_def tape_of_nat_def) next show "\\tap. tap = ([], r)\ tm_erase_right_then_dblBk_left \inv_tm_erase_right_then_dblBk_left_dnp_s0 r \" proof (rule Hoare_haltI) fix l::"cell list" fix r'':: "cell list" assume major: "(l, r'') = ([], r)" show "\stp. is_final (steps0 (1, l, r'') tm_erase_right_then_dblBk_left stp) \ inv_tm_erase_right_then_dblBk_left_dnp_s0 r holds_for steps0 (1, l, r'') tm_erase_right_then_dblBk_left stp" proof - from major and assms have "\stp. is_final (steps0 (1, l, r'') tm_erase_right_then_dblBk_left stp)" by auto then obtain stp where w_stp: "is_final (steps0 (1, l, r'') tm_erase_right_then_dblBk_left stp)" by blast moreover have "inv_tm_erase_right_then_dblBk_left_dnp_s0 r holds_for steps0 (1, l, r'') tm_erase_right_then_dblBk_left stp" proof - have "inv_tm_erase_right_then_dblBk_left_dnp r (1, l, r'')" by (simp add: major tape_of_list_def tape_of_nat_def) then have "inv_tm_erase_right_then_dblBk_left_dnp r (steps0 (1, l, r'') tm_erase_right_then_dblBk_left stp)" using inv_tm_erase_right_then_dblBk_left_dnp_steps by auto then show ?thesis - by (smt holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_dnp.simps is_final_eq w_stp) + by (smt (verit) holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_dnp.simps is_final_eq w_stp) qed ultimately show ?thesis by auto qed qed qed (* Total correctness of tm_erase_right_then_dblBk_left for DO NOTHING path *) (* Assemble a lexicographic measure function for the DO NOTHING path *) definition measure_tm_erase_right_then_dblBk_left_dnp :: "(config \ config) set" where "measure_tm_erase_right_then_dblBk_left_dnp = measures [ \(s, l, r). (if s = 0 then 0 else 5 - s) ]" lemma wf_measure_tm_erase_right_then_dblBk_left_dnp: "wf measure_tm_erase_right_then_dblBk_left_dnp" unfolding measure_tm_erase_right_then_dblBk_left_dnp_def by (auto) lemma measure_tm_erase_right_then_dblBk_left_dnp_induct [case_names Step]: "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_tm_erase_right_then_dblBk_left_dnp\ \ \n. P (f n)" using wf_measure_tm_erase_right_then_dblBk_left_dnp by (metis wf_iff_no_infinite_down_chain) lemma tm_erase_right_then_dblBk_left_dnp_halts: "\stp. is_final (steps0 (1, [], r) tm_erase_right_then_dblBk_left stp)" proof (induct rule: measure_tm_erase_right_then_dblBk_left_dnp_induct) case (Step stp) then have not_final: "\ is_final (steps0 (1, [], r) tm_erase_right_then_dblBk_left stp)" . have INV: "inv_tm_erase_right_then_dblBk_left_dnp r (steps0 (1, [], r) tm_erase_right_then_dblBk_left stp)" proof (rule_tac inv_tm_erase_right_then_dblBk_left_dnp_steps) show "inv_tm_erase_right_then_dblBk_left_dnp r (1, [], r)" by (simp add: tape_of_list_def tape_of_nat_def ) qed have SUC_STEP_RED: "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = step0 (steps0 (1, [], r) tm_erase_right_then_dblBk_left stp) tm_erase_right_then_dblBk_left" by (rule step_red) show "( steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp), steps0 (1, [], r) tm_erase_right_then_dblBk_left stp ) \ measure_tm_erase_right_then_dblBk_left_dnp" proof (cases "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp") case (fields s l r2) then have cf_cases: "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (s, l, r2)" . show ?thesis proof (rule tm_erase_right_then_dblBk_left_dnp_cases) from INV and cf_cases show "inv_tm_erase_right_then_dblBk_left_dnp r (s, l, r2)" by auto next assume "s=0" (* not possible *) with cf_cases not_final show ?thesis by auto next assume "s=1" show ?thesis proof (cases r) case Nil then have "r = []" . from cf_cases and \s=1\ have "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (1, l, r2)" by auto with cf_cases and \s=1\ and INV have "l = [] \ r = r2" by auto with cf_cases and \s=1\ and SUC_STEP_RED have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = step0 (1, [], r) tm_erase_right_then_dblBk_left" by auto also with \r = []\ and \l = [] \ r = r2\ have "... = (2,[],Bk#r2)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = (2,[],Bk#r2)" by auto with \steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (1, l, r2)\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_dnp_def) next case (Cons a rs) then have "r = a # rs" . then show ?thesis proof (cases a) case Bk then have "a = Bk" . from cf_cases and \s=1\ have "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (1, l, r2)" by auto with cf_cases and \s=1\ and INV have "l = [] \ r = r2" by auto with cf_cases and \s=1\ and SUC_STEP_RED have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = step0 (1, [], r) tm_erase_right_then_dblBk_left" by auto also with \r = a # rs\ and \a=Bk\ and \l = [] \ r = r2\ have "... = (2,[],Bk#r2)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = (2,[],Bk#r2)" by auto with \steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (1, l, r2)\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_dnp_def) next case Oc then have "a = Oc" . from cf_cases and \s=1\ have "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (1, l, r2)" by auto with cf_cases and \s=1\ and INV have "l = [] \ r = r2" by auto with cf_cases and \s=1\ and SUC_STEP_RED have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = step0 (1, [], r) tm_erase_right_then_dblBk_left" by auto also with \r = a # rs\ and \a=Oc\ and \l = [] \ r = r2\ have "... = (2,[],Bk#r2)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = (2,[],Bk#r2)" by auto with \steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (1, l, r2)\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_dnp_def) qed qed next assume "s=2" with cf_cases have "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (2, l, r2)" by auto with cf_cases and \s=2\ and INV have "(l = [] \ r2 = Bk#r)" by auto with cf_cases and \s=2\ and SUC_STEP_RED have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = step0 (2, l, r2) tm_erase_right_then_dblBk_left" by auto also with \s=2\ and \(l = [] \ r2 = Bk#r)\ have "... = (3,[],Bk#r2)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = (3,[],Bk#r2)" by auto with \steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (2, l, r2)\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_dnp_def) next assume "s=3" with cf_cases have "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (3, l, r2)" by auto with cf_cases and \s=3\ and INV have "(l = [] \ r2 = Bk#Bk#r)" by auto with cf_cases and \s=3\ and SUC_STEP_RED have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = step0 (3, l, r2) tm_erase_right_then_dblBk_left" by auto also with \s=3\ and \(l = [] \ r2 = Bk#Bk#r)\ have "... = (4,[Bk],Bk#r)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = (4,[Bk],Bk#r)" by auto with \steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (3, l, r2)\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_dnp_def) next assume "s=4" with cf_cases have "steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (4, l, r2)" by auto with cf_cases and \s=4\ and INV have "(l = [Bk] \ r2 = Bk#r)" by auto with cf_cases and \s=4\ and SUC_STEP_RED have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = step0 (4, l, r2) tm_erase_right_then_dblBk_left" by auto also with \s=4\ and \(l = [Bk] \ r2 = Bk#r)\ have "... = (0,[Bk,Bk],r)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [], r) tm_erase_right_then_dblBk_left (Suc stp) = (0,[Bk,Bk],r)" by auto with \steps0 (1, [], r) tm_erase_right_then_dblBk_left stp = (4, l, r2)\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_dnp_def) qed qed qed lemma tm_erase_right_then_dblBk_left_dnp_total_correctness: "\ \tap. tap = ([], r ) \ tm_erase_right_then_dblBk_left \ \tap. tap = ([Bk,Bk], r ) \" proof (rule tm_erase_right_then_dblBk_left_dnp_partial_correctness) show "\stp. is_final (steps0 (1, [], r) tm_erase_right_then_dblBk_left stp)" using tm_erase_right_then_dblBk_left_dnp_halts by auto qed (* ----------------------------------------------------------------- *) (* ERASE path tm_erase_right_then_dblBk_left_erp *) (* Sequence of states on the ERASE path: 1,2,3 then 5... *) (* ----------------------------------------------------------------- *) (* * \ \tap. tap = (Bk# rev(), ) \ * tm_erase_right_then_dblBk_left * \ \tap. \rex. tap = ([], [Bk,Bk] @ () @ [Bk] @ Bk \ rex ) \" *) (* Definition of invariants for the ERASE path: had to split definitions *) fun inv_tm_erase_right_then_dblBk_left_erp_s1 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s1 CL CR (l, r) = (l = [Bk,Oc] @ CL \ r = CR)" fun inv_tm_erase_right_then_dblBk_left_erp_s2 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s2 CL CR (l, r) = (l = [Oc] @ CL \ r = Bk#CR)" fun inv_tm_erase_right_then_dblBk_left_erp_s3 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s3 CL CR (l, r) = (l = CL \ r = Oc#Bk#CR)" fun inv_tm_erase_right_then_dblBk_left_erp_s5 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s5 CL CR (l, r) = (l = [Oc] @ CL \ r = Bk#CR)" fun inv_tm_erase_right_then_dblBk_left_erp_s6 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s6 CL CR (l, r) = (l = [Bk,Oc] @ CL \ ( (CR = [] \ r = CR) \ (CR \ [] \ (r = CR \ r = Bk # tl CR)) ) )" (* ENHANCE: simplify invariant of s6 (simpler to use for proof by cases "inv_tm_erase_right_then_dblBk_left_erp_s6 CL CR (l, r) = (l = [Bk,Oc] @ CL \ CR = [] \ r = CR \ l = [Bk,Oc] @ CL \ CR \ [] \ r = Oc # tl CR \ l = [Bk,Oc] @ CL \ CR \ [] \ r = Bk # tl CR )" *) fun inv_tm_erase_right_then_dblBk_left_erp_s7 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s7 CL CR (l, r) = ((\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs. CR = rs @ r) )" fun inv_tm_erase_right_then_dblBk_left_erp_s8 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s8 CL CR (l, r) = ((\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs1 rs2. CR = rs1 @ [Oc] @ rs2 \ r = Bk#rs2) )" fun inv_tm_erase_right_then_dblBk_left_erp_s9 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s9 CL CR (l, r) = ((\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs. CR = rs @ [Bk] @ r \ CR = rs \ r = []) )" fun inv_tm_erase_right_then_dblBk_left_erp_s10 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s10 CL CR (l, r) = ( (\lex rex. l = Bk \ lex @ [Bk,Oc] @ CL \ r = Bk \ Suc rex) \ (\rex. l = [Oc] @ CL \ r = Bk \ Suc rex) \ (\rex. l = CL \ r = Oc # Bk \ Suc rex) )" fun inv_tm_erase_right_then_dblBk_left_erp_s11 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s11 CL CR (l, r) = ( (\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)) \ (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk ) \ (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc ) \ (\rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]) \ (\rex ls1 ls2. l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc] ) \ (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk] ) \ (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc] ) \ (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ []) )" (* YYYY1' (\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc) ) \ (s11 \ s11) YYYY2' (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk) \ (s11 \ s11) YYYY2'' (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc) \ (s11 \ s11) YYYY5 (\rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]) \ (s10 \ s11) YYYY6 (\rex ls1 ls2. l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc] ) \ (s10 \ s11) YYYY3 (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk] ) \ (s10 \ s11) YYYY7 (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc] ) \ (s10 \ s11) YYYY8 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ []) (s11 \ s11) *) fun inv_tm_erase_right_then_dblBk_left_erp_s12 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s12 CL CR (l, r) = ( (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc) \ (\rex. l = [] \ r = Bk#rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk) \ (\rex. l = [] \ r = Bk#Bk#rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc) ) \ False )" (* YYYY4 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc) YYYY6''' (\rex. l = [] \ r = Bk#rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk) YYYY9 (\rex. l = [] \ r = Bk#Bk#rev CL @ Oc # Bk \ Suc rex) \ *) fun inv_tm_erase_right_then_dblBk_left_erp_s0 :: "(cell list) \ (cell list) \ tape \ bool" where "inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR (l, r) = ( (\rex. l = [] \ r = [Bk, Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex \ (CL = [] \ last CL = Oc) ) \ (\rex. l = [] \ r = [Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex \ CL \ [] \ last CL = Bk ) )" fun inv_tm_erase_right_then_dblBk_left_erp :: "(cell list) \ (cell list) \ config \ bool" where "inv_tm_erase_right_then_dblBk_left_erp CL CR (s, tap) = (if s = 0 then inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR tap else if s = 1 then inv_tm_erase_right_then_dblBk_left_erp_s1 CL CR tap else if s = 2 then inv_tm_erase_right_then_dblBk_left_erp_s2 CL CR tap else if s = 3 then inv_tm_erase_right_then_dblBk_left_erp_s3 CL CR tap else if s = 5 then inv_tm_erase_right_then_dblBk_left_erp_s5 CL CR tap else if s = 6 then inv_tm_erase_right_then_dblBk_left_erp_s6 CL CR tap else if s = 7 then inv_tm_erase_right_then_dblBk_left_erp_s7 CL CR tap else if s = 8 then inv_tm_erase_right_then_dblBk_left_erp_s8 CL CR tap else if s = 9 then inv_tm_erase_right_then_dblBk_left_erp_s9 CL CR tap else if s = 10 then inv_tm_erase_right_then_dblBk_left_erp_s10 CL CR tap else if s = 11 then inv_tm_erase_right_then_dblBk_left_erp_s11 CL CR tap else if s = 12 then inv_tm_erase_right_then_dblBk_left_erp_s12 CL CR tap else False)" lemma tm_erase_right_then_dblBk_left_erp_cases: fixes s::nat assumes "inv_tm_erase_right_then_dblBk_left_erp CL CR (s,l,r)" and "s=0 \ P" and "s=1 \ P" and "s=2 \ P" and "s=3 \ P" and "s=5 \ P" and "s=6 \ P" and "s=7 \ P" and "s=8 \ P" and "s=9 \ P" and "s=10 \ P" and "s=11 \ P" and "s=12 \ P" shows "P" proof - have "s < 4 \ 4 < s \ s < 13" proof (rule ccontr) assume "\ (s < 4 \ 4 < s \ s < 13)" with \inv_tm_erase_right_then_dblBk_left_erp CL CR (s,l,r)\ show False by auto qed then have "s = 0 \ s = 1 \ s = 2 \ s = 3 \ s = 5 \ s = 6 \ s = 7 \ s = 8 \ s = 9 \ s = 10 \ s = 11 \ s = 12" by arith with assms show ?thesis by auto qed (* -------------- step - lemma for the invariants of the ERASE path of tm_erase_right_then_dblBk_left ------------------ *) lemma inv_tm_erase_right_then_dblBk_left_erp_step: assumes "inv_tm_erase_right_then_dblBk_left_erp CL CR cf" and "noDblBk CL" and "noDblBk CR" shows "inv_tm_erase_right_then_dblBk_left_erp CL CR (step0 cf tm_erase_right_then_dblBk_left)" proof (cases cf) case (fields s l r) then have cf_cases: "cf = (s, l, r)" . show "inv_tm_erase_right_then_dblBk_left_erp CL CR (step0 cf tm_erase_right_then_dblBk_left)" proof (rule tm_erase_right_then_dblBk_left_erp_cases) from cf_cases and assms show "inv_tm_erase_right_then_dblBk_left_erp CL CR (s, l, r)" by auto next assume "s = 1" with cf_cases and assms have "(l = [Bk,Oc] @ CL \ r = CR)" by auto show ?thesis proof (cases r) case Nil then have "r = []" . with assms and cf_cases and \s = 1\ show ?thesis by (auto simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps) next case (Cons a rs) then have "r = a # rs" . show ?thesis proof (cases a) case Bk with assms and cf_cases and \r = a # rs\ and \s = 1\ show ?thesis by (auto simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps) next case Oc with assms and cf_cases and \r = a # rs\ and \s = 1\ show ?thesis by (auto simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps) qed qed next assume "s = 2" with cf_cases and assms have "l = [Oc] @ CL \ r = Bk#CR" by auto then have "step0 (2, [Oc] @ CL, Bk#CR) tm_erase_right_then_dblBk_left = (3, CL, Oc# Bk # CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (3, CL, Oc# Bk # CR)" by auto ultimately show ?thesis using assms and cf_cases and \s = 2\ and \l = [Oc] @ CL \ r = Bk#CR\ by auto next assume "s = 3" with cf_cases and assms have "l = CL \ r = Oc#Bk#CR" by auto then have "step0 (3, CL, Oc#Bk#CR) tm_erase_right_then_dblBk_left = (5, Oc#CL, Bk # CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (5, Oc#CL, Bk # CR)" by auto ultimately show ?thesis using assms and cf_cases and \s = 3\ and \l = CL \ r = Oc#Bk#CR\ by auto next assume "s = 5" with cf_cases and assms have "l = [Oc] @ CL \ r = Bk#CR" by auto then have "step0 (5, [Oc] @ CL, Bk#CR) tm_erase_right_then_dblBk_left = (6, Bk#Oc#CL, CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (6, Bk#Oc#CL, CR)" proof (cases CR) case Nil then show ?thesis by auto next case (Cons a cs) then have "CR = a # cs" . with \l = [Oc] @ CL \ r = Bk#CR\ and \s = 5\ and \CR = a # cs\ have "inv_tm_erase_right_then_dblBk_left_erp_s6 CL CR (Bk#Oc#CL, CR)" by simp with \s=5\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (6, Bk#Oc#CL, CR)" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 5\ and \l = [Oc] @ CL \ r = Bk#CR\ by auto next assume "s = 6" with cf_cases and assms have "l = [Bk,Oc] @ CL" and "( (CR = [] \ r = CR) \ (CR \ [] \ (r = CR \ r = Bk # tl CR)) )" by auto from \( (CR = [] \ r = CR) \ (CR \ [] \ (r = CR \ r = Bk # tl CR)) )\ show ?thesis proof assume "CR = [] \ r = CR" have "step0 (6, [Bk,Oc] @ CL, []) tm_erase_right_then_dblBk_left = (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, [])" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, [])" by auto ultimately show ?thesis using assms and cf_cases and \s = 6\ and \l = [Bk,Oc] @ CL\ and \CR = [] \ r = CR\ by auto next assume "CR \ [] \ (r = CR \ r = Bk # tl CR)" then have "CR \ []" and "r = CR \ r = Bk # tl CR" by auto from \r = CR \ r = Bk # tl CR\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (step0 cf tm_erase_right_then_dblBk_left)" proof assume "r = CR" show "inv_tm_erase_right_then_dblBk_left_erp CL CR (step0 cf tm_erase_right_then_dblBk_left)" proof (cases r) case Nil then have "r = []" . then have "step0 (6, [Bk,Oc] @ CL, []) tm_erase_right_then_dblBk_left = (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, [])" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, [])" by auto ultimately show ?thesis using assms and cf_cases and \s = 6\ and \l = [Bk,Oc] @ CL\ and \r = []\ by auto next case (Cons a rs') then have "r = a # rs'" . with \r = CR\ have "r = a # tl CR" by auto show ?thesis proof (cases a) case Bk then have "a = Bk" . then have "step0 (6, [Bk,Oc] @ CL, Bk # tl CR) tm_erase_right_then_dblBk_left = (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, tl CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, tl CR)" proof - from \CR \ []\ and \r = CR\ and \a = Bk\ and \r = a # tl CR\ and \l = [Bk,Oc] @ CL\ have "inv_tm_erase_right_then_dblBk_left_erp_s7 CL CR (Bk \ Suc 0 @ [Bk,Oc] @ CL, tl CR)" by (metis append.left_neutral append_Cons empty_replicate inv_tm_erase_right_then_dblBk_left_erp_s7.simps replicate_Suc ) then show "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, tl CR)" by auto qed ultimately show ?thesis using \CR \ []\ and \r = CR\ and \a = Bk\ and \r = a # tl CR\ and \l = [Bk,Oc] @ CL\ and \s = 6\ and cf_cases by auto next case Oc then have "a = Oc" . then have "step0 (6, [Bk,Oc] @ CL, Oc # rs') tm_erase_right_then_dblBk_left = (6, [Bk,Oc] @ CL, Bk # rs')" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (6, [Bk,Oc] @ CL, Bk # rs')" proof - from \CR \ []\ and \r = CR\ and \a = Oc\ and \r = a # tl CR\ and \l = [Bk,Oc] @ CL\ have "inv_tm_erase_right_then_dblBk_left_erp_s6 CL CR ([Bk,Oc] @ CL, Bk # rs')" using inv_tm_erase_right_then_dblBk_left_erp_s6.simps list.sel(3) local.Cons by blast then show "inv_tm_erase_right_then_dblBk_left_erp CL CR (6, [Bk,Oc] @ CL, Bk # rs')" by auto qed ultimately show ?thesis using \CR \ []\ and \r = CR\ and \a = Oc\ and \r = a # tl CR\ and \l = [Bk,Oc] @ CL\ and \s = 6\ and cf_cases by auto qed qed next assume "r = Bk # tl CR" have "step0 (6, [Bk,Oc] @ CL, Bk # tl CR) tm_erase_right_then_dblBk_left = (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, tl CR)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with \CR \ []\ and \r = Bk # tl CR\ have "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, tl CR)" proof - have "(\lex. Bk \ Suc 0 @ [Bk,Oc] @ CL = Bk \ Suc lex @ [Bk,Oc] @ CL) " by blast moreover with \CR \ []\ have "(\rs. CR = rs @ tl CR)" by (metis append_Cons append_Nil list.exhaust list.sel(3)) ultimately show "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc 0 @ [Bk,Oc] @ CL, tl CR)" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 6\ and \CR \ []\ and \r = Bk # tl CR\ and \l = [Bk,Oc] @ CL\ and cf_cases by auto qed qed next assume "s = 7" with cf_cases and assms have "(\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs. CR = rs @ r)" by auto then obtain lex rs where w_lex_rs: "l = Bk \ Suc lex @ [Bk,Oc] @ CL \ CR = rs @ r" by blast show ?thesis proof (cases r) case Nil then have "r=[]" . with w_lex_rs have "CR = rs" by auto have "step0 (7, Bk \ Suc lex @ [Bk,Oc] @ CL, []) tm_erase_right_then_dblBk_left = (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, [])" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with \CR = rs\ have "inv_tm_erase_right_then_dblBk_left_erp CL CR (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, [])" proof - have "(\lex'. Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL = Bk \ Suc lex' @ [Bk,Oc] @ CL)" by blast moreover have "\rs. CR = rs" by auto ultimately show "inv_tm_erase_right_then_dblBk_left_erp CL CR (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, [])" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 7\ and w_lex_rs and \CR = rs\ and \r=[]\ by auto next case (Cons a rs') then have "r = a # rs'" . show ?thesis proof (cases a) case Bk then have "a = Bk" . with w_lex_rs and \r = a # rs'\ have "CR = rs@(Bk#rs')" by auto have "step0 (7, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs') tm_erase_right_then_dblBk_left = (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs')" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with \CR = rs@(Bk#rs')\ have "inv_tm_erase_right_then_dblBk_left_erp CL CR (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs')" proof - have "(\lex'. Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL = Bk \ Suc lex' @ [Bk,Oc] @ CL)" by blast moreover with \r = a # rs'\ and \a = Bk\ and \CR = rs@(Bk#rs')\ have "\rs. CR = rs @ [Bk] @ rs'" by auto ultimately show "inv_tm_erase_right_then_dblBk_left_erp CL CR (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs')" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 7\ and w_lex_rs and \a = Bk\ and \r = a # rs'\ by simp next case Oc then have "a = Oc" . with w_lex_rs and \r = a # rs'\ have "CR = rs@(Oc#rs')" by auto have "step0 (7, Bk \ Suc lex @ [Bk,Oc] @ CL, Oc#rs') tm_erase_right_then_dblBk_left = (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" proof - have "(\lex'. Bk \ Suc lex @ [Bk,Oc] @ CL = Bk \ Suc lex' @ [Bk,Oc] @ CL)" by blast moreover with \r = a # rs'\ and \a = Oc\ and \CR = rs@(Oc#rs')\ have "\rs1 rs2. CR = rs1 @ [Oc] @ rs2 \ Bk#rs' = Bk#rs2" by auto ultimately show "inv_tm_erase_right_then_dblBk_left_erp CL CR (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 7\ and w_lex_rs and \a = Oc\ and \r = a # rs'\ by simp qed qed next assume "s = 8" with cf_cases and assms have "((\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs1 rs2. CR = rs1 @ [Oc] @ rs2 \ r = Bk#rs2) )" by auto then obtain lex rs1 rs2 where w_lex_rs1_rs2: "l = Bk \ Suc lex @ [Bk,Oc] @ CL \ CR = rs1 @ [Oc] @ rs2 \ r = Bk#rs2" by blast have "step0 (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs2) tm_erase_right_then_dblBk_left = (7, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs2)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs2)" proof - have "(\lex'. Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL = Bk \ Suc lex' @ [Bk,Oc] @ CL)" by blast moreover have "\rs. CR = rs @ []" by auto ultimately show "inv_tm_erase_right_then_dblBk_left_erp CL CR (7, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs2)" using w_lex_rs1_rs2 by auto qed ultimately show ?thesis using assms and cf_cases and \s = 8\ and w_lex_rs1_rs2 by auto next assume "s = 9" with cf_cases and assms have "(\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs. CR = rs @ [Bk] @ r \ CR = rs \ r = [])" by auto then obtain lex rs where w_lex_rs: "l = Bk \ Suc lex @ [Bk,Oc] @ CL \ (CR = rs @ [Bk] @ r \ CR = rs \ r = [])" by blast then have "CR = rs @ [Bk] @ r \ CR = rs \ r = []" by auto then show ?thesis proof assume "CR = rs \ r = []" have "step0 (9, Bk \ Suc lex @ [Bk,Oc] @ CL, []) tm_erase_right_then_dblBk_left = (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with \CR = rs \ r = []\ have "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" proof - from w_lex_rs and \CR = rs \ r = []\ have "\lex' rex. Bk \ lex @ [Bk,Oc] @ CL = Bk \ lex' @ [Bk,Oc] @ CL \ [Bk] = Bk \ Suc rex" by (simp) then show "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 9\ and w_lex_rs and \CR = rs \ r = []\ by auto next assume "CR = rs @ [Bk] @ r" show ?thesis proof (cases r) case Nil then have "r=[]" . have "step0 (9, Bk \ Suc lex @ [Bk,Oc] @ CL, []) tm_erase_right_then_dblBk_left = (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with \CR = rs @ [Bk] @ r\ have "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" proof - from w_lex_rs and \CR = rs @ [Bk] @ r\ have "\lex' rex. Bk \ lex @ [Bk,Oc] @ CL = Bk \ lex' @ [Bk,Oc] @ CL \ [Bk] = Bk \ Suc rex" by (simp) with \s=9\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 9\ and w_lex_rs and \r=[]\ by auto next case (Cons a rs') then have "r = a # rs'" . show ?thesis proof (cases a) case Bk then have "a = Bk" . with \CR = rs @ [Bk] @ r\ and \r = a # rs'\ have "CR = rs @ [Bk] @ Bk # rs'" by auto moreover from assms have "noDblBk CR" by auto ultimately have False using hasDblBk_L1 by auto then show ?thesis by auto next case Oc then have "a = Oc" . with \CR = rs @ [Bk] @ r\ and \r = a # rs'\ have "CR = rs @ [Bk] @ Oc # rs'" by auto have "step0 (9, Bk \ Suc lex @ [Bk,Oc] @ CL, Oc # rs') tm_erase_right_then_dblBk_left = (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk # rs')" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk # rs')" proof - have "(\lex'. Bk \ Suc lex @ [Bk,Oc] @ CL = Bk \ Suc lex' @ [Bk,Oc] @ CL)" by blast moreover with \r = a # rs'\ and \a = Oc\ and \CR = rs @ [Bk] @ Oc # rs'\ have "\rs1 rs2. CR = rs1 @ [Oc] @ rs2 \ Bk#rs' = Bk#rs2" by auto ultimately show "inv_tm_erase_right_then_dblBk_left_erp CL CR (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 9\ and w_lex_rs and \a = Oc\ and \r = a # rs'\ by simp qed qed qed next assume "s = 10" with cf_cases and assms have "(\lex rex. l = Bk \ lex @ [Bk,Oc] @ CL \ r = Bk \ Suc rex) \ (\rex. l = [Oc] @ CL \ r = Bk \ Suc rex) \ (\rex. l = CL \ r = Oc # Bk \ Suc rex)" by auto then obtain lex rex where w_lex_rex: "l = Bk \ lex @ [Bk,Oc] @ CL \ r = Bk \ Suc rex \ l = [Oc] @ CL \ r = Bk \ Suc rex \ l = CL \ r = Oc # Bk \ Suc rex" by blast then show ?thesis proof assume "l = Bk \ lex @ [Bk, Oc] @ CL \ r = Bk \ Suc rex" then have "l = Bk \ lex @ [Bk, Oc] @ CL" and "r = Bk \ Suc rex" by auto show ?thesis proof (cases lex) case 0 with \l = Bk \ lex @ [Bk, Oc] @ CL\ have "l = [Bk, Oc] @ CL" by auto have "step0 (10, [Bk, Oc] @ CL, Bk \ Suc rex) tm_erase_right_then_dblBk_left = (10, [Oc] @ CL, Bk \ Suc (Suc rex))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, [Oc] @ CL, Bk \ Suc (Suc rex))" proof - from \l = [Bk, Oc] @ CL\ and \r = Bk \ Suc rex\ have "\rex'. [Oc] @ CL = [Oc] @ CL \ Bk \ Suc (Suc rex) = Bk \ Suc rex'" by blast with \l = [Bk, Oc] @ CL\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, [Oc] @ CL, Bk \ Suc (Suc rex))" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 10\ and \l = [Bk, Oc] @ CL\ and \r = Bk \ Suc rex\ by auto next case (Suc nat) then have "lex = Suc nat" . with \l = Bk \ lex @ [Bk, Oc] @ CL\ have "l= Bk \ Suc nat @ [Bk, Oc] @ CL" by auto have "step0 (10, Bk \ Suc nat @ [Bk, Oc] @ CL, Bk \ Suc rex) tm_erase_right_then_dblBk_left = (10, Bk \ nat @ [Bk, Oc] @ CL, Bk \ Suc (Suc rex))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, Bk \ nat @ [Bk, Oc] @ CL, Bk \ Suc (Suc rex))" proof - from \l= Bk \ Suc nat @ [Bk, Oc] @ CL\ and \r = Bk \ Suc rex\ have "\lex' rex'. Bk \ Suc nat @ [Bk, Oc] @ CL = Bk \ lex' @ [Bk,Oc] @ CL \ Bk \ Suc (Suc rex) = Bk \ Suc rex'" by blast with \l= Bk \ Suc nat @ [Bk, Oc] @ CL\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, Bk \ nat @ [Bk, Oc] @ CL, Bk \ Suc (Suc rex))" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 10\ and \l= Bk \ Suc nat @ [Bk, Oc] @ CL\ and \r = Bk \ Suc rex\ by auto qed next assume "l = [Oc] @ CL \ r = Bk \ Suc rex \ l = CL \ r = Oc # Bk \ Suc rex" then show ?thesis proof assume "l = [Oc] @ CL \ r = Bk \ Suc rex" then have "l = [Oc] @ CL" and "r = Bk \ Suc rex" by auto have "step0 (10, [Oc] @ CL, Bk \ Suc rex) tm_erase_right_then_dblBk_left = (10, CL, Oc# Bk \ (Suc rex))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, CL, Oc# Bk \ (Suc rex))" proof - from \l = [Oc] @ CL\ and \r = Bk \ Suc rex\ have "\rex'. [Oc] @ CL = [Oc] @ CL \ Bk \ Suc rex = Bk \ Suc rex'" by blast with \l = [Oc] @ CL\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (10, CL, Oc# Bk \ (Suc rex))" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 10\ and \l = [Oc] @ CL\ and \r = Bk \ Suc rex\ by auto next assume "l = CL \ r = Oc # Bk \ Suc rex" then have "l = CL" and "r = Oc # Bk \ Suc rex" by auto show ?thesis proof (cases CL) (* here, we start decomposing CL in the loop 'move to left until Oc *) case Nil then have "CL = []" . with \l = CL\ have "l = []" by auto have "step0 (10, [], Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, [], Bk#Oc# Bk \ (Suc rex))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#Oc# Bk \ (Suc rex))" proof - (* Special case: CL = [] from call (1, [Oc, Bk], []) we get 11: [] _11_ [Bk,Oc,Bk,Bk,Bk] YYYY1' "\rex'. ] = [] \ Bk# Oc# Bk \ Suc rex = [Bk] @ rev CL @ Oc # Bk \ Suc rex' \ (CL = [] \ last CL = Oc)" *) from \CL = []\ have "\rex'. [] = [] \ Bk# Oc# Bk \ Suc rex = [Bk] @ rev CL @ Oc # Bk \ Suc rex' \ (CL = [] \ last CL = Oc)" by auto with \l = []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#Oc# Bk \ (Suc rex))" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 10\ and \l = []\ and \r = Oc # Bk \ Suc rex\ by auto next case (Cons a cls) then have "CL = a # cls" . with \l = CL\ have "l = a # cls" by auto then show ?thesis proof (cases a) case Bk then have "a = Bk" . with \l = a # cls\ have "l = Bk # cls" by auto with \a = Bk\ \CL = a # cls\ have "CL = Bk # cls" by auto have "step0 (10, Bk # cls, Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, cls, Bk# Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Bk# Oc # Bk \ Suc rex)" proof (cases cls) case Nil then have "cls = []" . with \CL = Bk # cls\ have "CL = [Bk]" by auto (* Special case: CL ends with a blank from call ctxrunTM tm_erase_right_then_dblBk_left (1, [Bk, Oc,Bk], [Bk,Oc,Oc,Bk,Oc,Oc]) 24: [Bk] _10_ [Oc,Bk,Bk,Bk,Bk,Bk,Bk,Bk,Bk] 25: [] _11_ [Bk,Oc,Bk,Bk,Bk,Bk,Bk,Bk,Bk,Bk] YYYY2' "\rex'. [] = [] \ Bk# Oc# Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ \ CL \ [] \ last CL = Bk" *) then have "\rex'. [] = [] \ Bk# Oc# Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ CL \ [] \ last CL = Bk" by auto with \l = Bk # cls\ and \cls = []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Bk# Oc # Bk \ Suc rex)" by auto next case (Cons a cls') then have "cls = a # cls'" . then show ?thesis proof (cases a) case Bk with \CL = Bk # cls\ and \cls = a # cls'\ have "CL = Bk# Bk# cls'" by auto with \noDblBk CL\ have False using noDblBk_def by auto then show ?thesis by auto next case Oc then have "a = Oc" . with \CL = Bk # cls\ and \cls = a # cls'\ and \l = Bk # cls\ have "CL = Bk# Oc# cls' \ l = Bk # Oc # cls'" by auto (* from call (1, [Oc,Oc,Bk,Oc,Bk], [Oc,Oc,Oc,Bk,Oc,Oc]) we get 26: [Oc,Oc] _11_ [Bk,Oc,Bk,Bk,Bk,Bk,Bk,Bk,Bk,Bk] YYYY3 "\rex' ls1 ls2. Oc#cls' = Oc#ls2 \ Bk# Oc# Bk \ Suc rex = rev ls1 @ Oc # Bk \ Suc rex' \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk]" *) with \cls = a # cls'\ and \a=Oc\ have "\rex' ls1 ls2. Oc#cls' = Oc#ls2 \ Bk# Oc# Bk \ Suc rex = rev ls1 @ Oc # Bk \ Suc rex' \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk]" by auto with \cls = a # cls'\ and \a=Oc\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Bk# Oc # Bk \ Suc rex)" by auto qed qed ultimately show ?thesis using assms and cf_cases and \s = 10\ and \l = Bk # cls\ and \r = Oc # Bk \ Suc rex\ by auto next case Oc then have "a = Oc" . with \l = a # cls\ have "l = Oc # cls" by auto with \a = Oc\ \CL = a # cls\ have "CL = Oc # cls" by auto (* a normal case *) have "step0 (10, Oc # cls, Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, cls, Oc # Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Oc # Oc # Bk \ Suc rex)" proof (cases cls) case Nil then have "cls = []" . with \CL = Oc # cls\ have "CL = [Oc]" by auto (* from call (1, [Oc,Oc,Bk], [Oc,Oc,Oc,Bk,Oc,Oc]) we get 26: [] _11_ [Oc,Oc,Bk,Bk,Bk,Bk,Bk,Bk,Bk,Bk] YYYY2'' "\rex'. [] = [] \ Oc# Oc# Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ CL \ [] \ last CL = Oc" *) then have "\rex'. [] = [] \ Oc# Oc# Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ CL \ [] \ last CL = Oc" by auto with \l = Oc # cls\ and \cls = []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Oc# Oc # Bk \ Suc rex)" by auto next case (Cons a cls') then have "cls = a # cls'" . then show ?thesis proof (cases a) case Bk then have "a=Bk" . with \CL = Oc # cls\ and \cls = a # cls'\ and \l = Oc # cls\ have "CL = Oc# Bk# cls'" and "l = Oc # Bk # cls'" and "CL = l" and \cls = Bk # cls'\ by auto from \CL = Oc# Bk# cls'\ and \noDblBk CL\ have "cls' = [] \ (\cls''. cls' = Oc# cls'')" by (metis (full_types) \CL = Oc # cls\ \cls = Bk # cls'\ append_Cons append_Nil cell.exhaust hasDblBk_L1 neq_Nil_conv) then show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Oc # Oc # Bk \ Suc rex)" proof assume "cls' = []" with \CL = Oc# Bk# cls'\ and \CL = l\ and \cls = Bk # cls'\ have "CL = Oc# Bk# []" and "l = Oc# Bk# []" and "cls = [Bk]" by auto (* from call (1, [Bk,Oc,Oc,Bk], [Oc,Oc,Oc,Bk,Oc,Oc]) we get 26: [Bk] _11_ [Oc,Oc,Bk,Bk,Bk,Bk,Bk,Bk,Bk,Bk] YYYY5 "\rex'. cls = [Bk] \ Oc# Oc# Bk \ Suc rex = rev [Oc] @ Oc # Bk \ Suc rex' \ CL = [Oc, Bk]" *) then have "\rex'. cls = [Bk] \ Oc# Oc# Bk \ Suc rex = rev [Oc] @ Oc # Bk \ Suc rex' \ CL = [Oc, Bk]" by auto with \CL = Oc# Bk# []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Oc # Oc # Bk \ Suc rex)" by auto next assume "\cls''. cls' = Oc # cls''" then obtain cls'' where "cls' = Oc # cls''" by blast with \CL = Oc# Bk# cls'\ and \CL = l\ and \cls = Bk # cls'\ have "CL = Oc# Bk# Oc # cls''" and "l = Oc# Bk# Oc # cls''" and "cls = Bk#Oc # cls''" by auto (* very special case call (1, [*,Oc,Bk, Oc,Oc,Bk], [Oc,Oc,Oc,Bk,Oc,Oc]) trailing Bk on initial left tape from call (1, [Oc,Bk, Oc,Oc,Bk], [Oc,Oc,Oc,Bk,Oc,Oc]) we get 26: [Oc,Bk] _11_ [Oc,Oc,Bk,Bk,Bk,Bk,Bk,Bk,Bk,Bk] from call (1, [Oc,Oc,Bk,Oc,Bk, Oc,Oc,Bk], [Oc,Oc,Oc,Bk,Oc,Oc]) we get 26: [Oc,Oc,Bk,Oc,Bk] _11_ [Oc,Oc,Bk,Bk,Bk,Bk,Bk,Bk,Bk,Bk] YYYY6 "\rex' ls1 ls2. cls = Bk#Oc#ls2 \ Oc# Oc# Bk \ Suc rex = rev ls1 @ Oc # Bk \ Suc rex' \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc]" *) then have "\rex' ls1 ls2. cls = Bk#Oc#ls2 \ Oc# Oc# Bk \ Suc rex = rev ls1 @ Oc # Bk \ Suc rex' \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc]" by auto then show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Oc # Oc # Bk \ Suc rex)" by auto qed next case Oc then have "a=Oc" . with \CL = Oc # cls\ and \cls = a # cls'\ and \l = Oc # cls\ have "CL = Oc# Oc# cls'" and "l = Oc # Oc # cls'" and "CL = l" and \cls = Oc # cls'\ by auto (* We know more : ls1 = [Oc] YYYY7 "\rex' ls1 ls2. cls = Oc#ls2 \ Oc# Oc# Bk \ Suc rex = rev ls1 @ Oc # Bk \ Suc rex' \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc]" "\rex' ls1 ls2. cls = Oc#ls2 \ Oc# Oc# Bk \ Suc rex = rev ls1 @ Oc # Bk \ Suc rex' \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc]" *) then have "\rex' ls1 ls2. cls = Oc#ls2 \ Oc# Oc# Bk \ Suc rex = rev ls1 @ Oc # Bk \ Suc rex' \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc]" by auto then show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, cls, Oc # Oc # Bk \ Suc rex)" by auto qed qed ultimately show ?thesis using assms and cf_cases and \s = 10\ and \l = Oc # cls\ and \r = Oc # Bk \ Suc rex\ by auto qed qed qed qed next assume "s = 11" with cf_cases and assms have "(\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)) \ (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk ) \ (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc ) \ (\rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]) \ (\rex ls1 ls2. l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc] ) \ (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk] ) \ (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc] ) \ (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [])" by auto then have s11_cases: "\P. \ \rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc) \ P; \rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk \ P; \rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc \ P; \rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk] \ P; \rex ls1 ls2. l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc] \ P; \rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk] \ P; \rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc] \ P; \rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ P \ \ P" by blast show ?thesis proof (rule s11_cases) assume "\rex ls1 ls2. l = Bk # Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk # Oc # ls2 \ ls1 = [Oc]" then obtain rex ls1 ls2 where A_case: "l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc]" by blast then have "step0 (11, Bk # Oc # ls2, r) tm_erase_right_then_dblBk_left = (11, Oc # ls2, Bk # r)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, Oc # ls2, Bk # r)" proof - from A_case (* YYYY8 *) have "\rex' ls1' ls2'. Oc#ls2 = ls2' \ Bk# Oc# Oc# Bk \ Suc rex' = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ tl ls1' \ []" by force with A_case show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, Oc # ls2, Bk # r)" by auto qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case by simp next assume "\rex ls1 ls2. l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Oc]" then obtain rex ls1 ls2 where A_case: "l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Oc]" by blast then have "step0 (11, Oc # ls2, r) tm_erase_right_then_dblBk_left = (11, ls2, Oc#r)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, ls2, Oc#r)" proof (rule noDblBk_cases) from \noDblBk CL\ show "noDblBk CL" . next from A_case show "CL = [Oc,Oc] @ ls2" by auto next assume "ls2 = []" (* with A_case (* YYYY7''' *) have "\rex'. ls2 = [] \ [Oc,Oc] @ Oc # Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ CL = [Oc, Oc]" by force *) with A_case (* YYYY2'' *) have "\rex'. ls2 = [] \ [Oc,Oc] @ Oc # Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex'" by force with A_case and \ls2 = []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, ls2, Oc#r)" by auto next assume "ls2 = [Bk]" (* YYYY8 *) with A_case have "\rex' ls1' ls2'. ls2 = ls2' \ Oc#Oc# Oc# Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ []" by simp with A_case and \ls2 = [Bk]\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, ls2, Oc#r)" by force next fix C3 assume "ls2 = Bk # Oc # C3" (* YYYY8 *) with A_case have "\rex' ls1' ls2'. ls2 = ls2' \ Oc#Oc# Oc# Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ []" by simp with A_case and \ls2 = Bk # Oc # C3\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, ls2, Oc#r)" by force next fix C3 assume "ls2 = Oc # C3" (* YYYY8 *) with A_case have "\rex' ls1' ls2'. ls2 = ls2' \ Oc#Oc# Oc# Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ []" by simp with A_case and \ls2 = Oc # C3\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, ls2, Oc#r)" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case by simp next assume "\rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]" then obtain rex where A_case: "l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]" by blast then have "step0 (11, [Bk] , r) tm_erase_right_then_dblBk_left = (11, [], Bk#r)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#r)" proof - from A_case (* have "\rex'. [] = [] \ Bk#rev [Oc] @ Oc # Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ CL = [Oc, Bk]" (* YYYY5' *) by simp *) have "\rex'. [] = [] \ Bk#rev [Oc] @ Oc # Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex'" (* YYYY2' *) by simp with A_case show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#r)" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case by simp next (* YYYY8 for s11 *) assume "\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ []" then obtain rex ls1 ls2 where A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ []" by blast then have "\z b bs. ls1 = z#bs@[b]" (* the symbol z does not matter *) by (metis Nil_tl list.exhaust_sel rev_exhaust) then have "\z bs. ls1 = z#bs@[Bk] \ ls1 = z#bs@[Oc]" using cell.exhaust by blast then obtain z bs where w_z_bs: "ls1 = z#bs@[Bk] \ ls1 = z#bs@[Oc]" by blast then show "inv_tm_erase_right_then_dblBk_left_erp CL CR (step0 cf tm_erase_right_then_dblBk_left)" proof assume major1: "ls1 = z # bs @ [Bk]" then have major2: "rev ls1 = Bk#(rev bs)@[z]" by auto (* in this case all transitions will be s11 \ s12 *) show ?thesis proof (rule noDblBk_cases) from \noDblBk CL\ show "noDblBk CL" . next from A_case show "CL = ls1 @ ls2" by auto next assume "ls2 = []" with A_case have "step0 (11, [] , Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (12, [], Bk#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, [], Bk#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" proof - from A_case \rev ls1 = Bk#(rev bs)@[z]\ and \ls2 = []\ have "ls1 = z#bs@[Bk]" and "CL = ls1" and "r = rev CL @ Oc # Bk \ Suc rex" by auto (* YYYY6''' \rex. l = [] \ r = Bk#rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk *) with A_case \rev ls1 = Bk#(rev bs)@[z]\ and \ls2 = []\ have "\rex'. [] = [] \ Bk#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex = Bk#rev CL @ Oc # Bk \ Suc rex' \ CL \ [] \ last CL = Bk" by simp with A_case \rev ls1 = Bk#(rev bs)@[z]\ and \ls2 = []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, [], Bk#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case and \rev ls1 = Bk#(rev bs)@[z]\ and \ls2 = []\ by simp next assume "ls2 = [Bk]" (* A_case: l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] *) with A_case \rev ls1 = Bk#(rev bs)@[z]\ and \ls2 = [Bk]\ have "ls1 = z#bs@[Bk]" and "CL = z#bs@[Bk]@[Bk]" by auto with \noDblBk CL\ have False by (metis A_case \ls2 = [Bk]\ append_Cons hasDblBk_L5 major2) then show ?thesis by auto next fix C3 assume minor: "ls2 = Bk # Oc # C3" with A_case and major2 have "CL = z # bs @ [Bk] @ Bk # Oc # C3" by auto with \noDblBk CL\ have False by (metis append.left_neutral append_Cons append_assoc hasDblBk_L1 major1 minor) then show ?thesis by auto next fix C3 assume minor: "ls2 = Oc # C3" (* A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ []" major1: "ls1 = Oc # bs @ [Bk]" major2: "rev ls1 = Bk # rev bs @ [Oc]" minor1: "ls2 = Oc # C3" l = Oc # C3 r = rev ls1 @ Oc # Bk \ Suc rex r = Bk#(rev bs)@[Oc] @ Oc # Bk \ Suc rex thus: s11 \ s12 (12, C3, Oc#Bk#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ) l' = C3 r' = Oc#Bk#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ls2' = C3 rev ls1' = Oc#Bk#(rev bs)@[Oc] ls1' = Oc# bs @ [Bk] @ [Oc] = ls1@[Oc] CL = ls1 @ ls2 = ls1 @ Oc # C3 = ls1 @ [Oc] @ [C3] = ls1' @ ls2' \ hd ls1 = Oc \ tl ls1 \ [] (\rex' ls1' ls2'. C3 = ls2' \ Oc#Bk#(rev bs)@[Oc] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ [] \ last ls1' = Oc) again YYYY4 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] \ last ls1 = Oc) *) with A_case have "step0 (11, Oc # C3 , Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (12, C3, Oc#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, C3, Oc#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" proof - from A_case and \rev ls1 = Bk # rev bs @ [z]\ and \ls2 = Oc # C3\ and \ls1 = z # bs @ [Bk]\ have "\rex' ls1' ls2'. C3 = ls2' \ Oc#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = z \ tl ls1' \ [] \ last ls1' = Oc" by simp with A_case \rev ls1 = Bk # rev bs @ [z]\ and \ls2 = Oc # C3\ and \ls1 = z # bs @ [Bk]\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, C3, Oc#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by simp qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case and \rev ls1 = Bk#(rev bs)@[z]\ and \ls2 = Oc # C3\ by simp qed next assume major1: "ls1 = z # bs @ [Oc]" then have major2: "rev ls1 = Oc#(rev bs)@[z]" by auto (* in this case all transitions will be s11 \ s11 *) show ?thesis proof (rule noDblBk_cases) from \noDblBk CL\ show "noDblBk CL" . next from A_case show "CL = ls1 @ ls2" by auto next assume "ls2 = []" (* A_case: l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] l = [] r = rev ls1 @ Oc # Bk \ Suc rex r = Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex l' = [] r' = Bk#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ls1 = Oc # bs @ [Oc] rev ls1 = Oc#(rev bs)@[Oc] rev ls1' = Bk#Oc#(rev bs)@[Oc] ls1' = Oc # bs @ [Oc] @ [Bk] ls2' = [] CL = ls1 @ ls2 = (Oc # bs @ [Oc]) @ [] = Oc # bs @ [Oc] = ls1 rev ls1 = rev CL \rex'. [] = [] \ Bk#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex = Bk#rev CL @ Oc # Bk \ Suc rex' \ last CL = Oc (\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ last CL = Oc) we simplify YYYY1' (\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex) now, we generalize to YYYY1' (\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex) \ (CL = [] \ last CL = Oc) *) with A_case have "step0 (11, [] , Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" proof - from A_case \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = []\ have "ls1 = z # bs @ [Oc]" and "CL = ls1" and "r = rev CL @ Oc # Bk \ Suc rex" by auto (* new invariant for s11: YYYY1' (\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)) *) with A_case \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = []\ have "\rex'. [] = [] \ Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex = Bk#rev CL @ Oc # Bk \ Suc rex' \ (CL = [] \ last CL = Oc)" by simp with A_case \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case and \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = []\ by simp next assume "ls2 = [Bk]" (* A_case: l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] l = [Bk] r = rev ls1 @ Oc # Bk \ Suc rex r = Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex l' = [] r' = Bk#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ls1 = Oc # bs @ [Oc] rev ls1 = Oc#(rev bs)@[Oc] rev ls1' = Bk#Oc#(rev bs)@[Oc] ls1' = Oc # bs @ [Oc] @ [Bk] = ls1 @ [Bk] ls2' = [] CL = ls1 @ ls2 = (Oc # bs @ [Oc]) @ [Bk] = ls1 @ [Bk] = ls1' list functions (hd ls) and (last ls) are only usefull if ls \ [] is known! new invariant for s11: YYYY2' (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk) \ (s11 \ s11) YYYY2'' (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc) \ (s11 \ s11) *) with A_case have "step0 (11, [Bk] , Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" proof - from A_case \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = [Bk]\ have "ls1 = z # bs @ [Oc]" and "CL = ls1@[Bk]" by auto (* new invariant for s11: YYYY2' (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk) *) with A_case \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = [Bk]\ have "\rex'. [] = [] \ Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ CL \ [] \ last CL = Bk" by simp with A_case \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = [Bk]\ and \CL = ls1@[Bk]\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case and \rev ls1 = Oc#(rev bs)@[z]\ and \ls2 = [Bk]\ by simp next fix C3 assume minor: "ls2 = Bk # Oc # C3" (* thm A_case: l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] thm major1: ls1 = Oc # bs @ [Oc] thm major2: rev ls1 = Oc # rev bs @ [Oc] minor1: "ls2 = Bk # Oc # C3" l = Bk # Oc # C3 r = Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex thus: s11 \ s11 (11, Oc # C3, Bk#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ) l' = Oc # C3 r' = Bk#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ls2' = Oc # C3 rev ls1' = Bk#Oc#(rev bs)@[Oc] ls1' = Oc# bs @ [Oc] @ [Bk] = ls1@[Bk] CL = ls1 @ ls2 = ls1 @ Bk # Oc # C3 = ls1 @ [Bk] @ [Oc ,C3] = ls1' @ ls2' \ hd ls1 = Oc \ tl ls1 \ [] (\rex' ls1' ls2'. Oc # C3 = ls2' \ Bk#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ []) again YYYY8 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ []) *) with A_case have "step0 (11, Bk # Oc # C3 , Oc # rev bs @ [z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, Oc # C3, Bk#Oc # rev bs @ [z] @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, Oc # C3, Bk#Oc # rev bs @ [z] @ Oc # Bk \ Suc rex )" proof - from A_case and \rev ls1 = Oc # rev bs @ [z]\ and \ls2 = Bk # Oc # C3\ and \ls1 = z # bs @ [Oc]\ have "\rex' ls1' ls2'. Oc # C3 = ls2' \ Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = z \ tl ls1' \ [] " by simp with A_case \rev ls1 = Oc # rev bs @ [z]\ and \ls2 = Bk # Oc # C3\ and \ls1 = z # bs @ [Oc]\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, Oc # C3, Bk#Oc # rev bs @ [z] @ Oc # Bk \ Suc rex )" by simp qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case and \rev ls1 = Oc # rev bs @ [z]\ and \ls2 = Bk # Oc # C3\ by simp next fix C3 assume minor: "ls2 = Oc # C3" (* A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ []" major1: "ls1 = Oc # bs @ [Oc]" major2: "rev ls1 = Oc # rev bs @ [Oc]" minor1: "ls2 = Oc # C3" l = Oc # C3 r = rev ls1 @ Oc # Bk \ Suc rex r = Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex thus: s11 \ s11 (11, C3, Oc#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ) l' = C3 r' = Oc#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex ls2' = C3 rev ls1' = Oc#Oc#(rev bs)@[Oc] ls1' = Oc# bs @ [Oc] @ [Oc] = ls1@[Oc] CL = ls1 @ ls2 = ls1 @ Oc # C3 = ls1 @ [Oc] @ [C3] = ls1' @ ls2' \ hd ls1 = Oc \ tl ls1 \ [] (\rex' ls1' ls2'. C3 = ls2' \ Oc#Oc#(rev bs)@[Oc] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ []) again YYYY8 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ []) *) with A_case have "step0 (11, Oc # C3 , Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, C3, Oc#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, C3, Oc#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" proof - from A_case and \rev ls1 = Oc # rev bs @ [z]\ and \ls2 = Oc # C3\ and \ls1 = z # bs @ [Oc]\ have "\rex' ls1' ls2'. C3 = ls2' \ Oc#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ hd ls1' = z \ tl ls1' \ []" by simp with A_case \rev ls1 = Oc # rev bs @ [z]\ and \ls2 = Oc # C3\ and \ls1 = z # bs @ [Oc]\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, C3, Oc#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by simp qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case and \rev ls1 = Oc # rev bs @ [z]\ and \ls2 = Oc # C3\ by simp qed qed next (* YYYY3 *) assume "\rex ls1 ls2. l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Bk]" then obtain rex ls1 ls2 where A_case: "l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Bk]" by blast then have major2: "rev ls1 = [Bk]" by auto (* "l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Bk]" l = Oc # ls2 r = rev ls1 @ Oc # Bk \ Suc rex r = [Bk] @ Oc # Bk \ Suc rex l' = ls2 r' = Oc#[Bk] @ Oc # Bk \ Suc rex \ s12: (12, ls2, Oc#[Bk] @ Oc # Bk \ Suc rex ) ls1 = [Bk] rev ls1 = [Bk] rev ls1' = Oc#[Bk] ls1' = Bk#[Oc] ls2' = ls2 CL = ls1 @ Oc # ls2 = [Bk] @ Oc # ls2 = (ls1 @ [Oc]) @ ls2 = ls1'@ls2' \ ls1' = Bk#[Oc] r' = Oc#[Bk] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex \rex' ls1' ls2'. ls2 = ls2' \ Oc#[Bk] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ ls1' = [Bk,Oc] YYYY4 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc) *) with A_case have "step0 (11, Oc # ls2 , [Bk] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (12, ls2, Oc#[Bk] @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, ls2, Oc#[Bk] @ Oc # Bk \ Suc rex )" proof - from A_case and \rev ls1 = [Bk]\ have "\rex' ls1' ls2'. ls2 = ls2' \ Oc#[Bk] @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ tl ls1' \ [] \ last ls1' = Oc" (* YYYY4 *) by simp with A_case \rev ls1 = [Bk]\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, ls2, Oc#[Bk] @ Oc # Bk \ Suc rex )" by simp qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case and \rev ls1 = [Bk]\ by simp next assume "\rex. l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" (* YYYY1' *) then obtain rex where A_case: "l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" by blast then have "step0 (11, [] , Bk # rev CL @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (12, [], Bk#Bk # rev CL @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, [], Bk#Bk # rev CL @ Oc # Bk \ Suc rex )" proof - from A_case have "\rex'. [] = [] \ Bk#Bk # rev CL @ Oc # Bk \ Suc rex = Bk# Bk # rev CL @ Oc # Bk \ Suc rex' \ (CL = [] \ last CL = Oc)" (* YYYY9 *) by simp with A_case show "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, [], Bk#Bk # rev CL @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case by simp next assume "\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" (* YYYY2' *) then obtain rex where A_case: "l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" by blast then have "hd (rev CL) = Bk" by (simp add: hd_rev) with A_case have "step0 (11, [] , rev CL @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (12, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, [], Bk # rev CL @ Oc # Bk \ Suc rex )" proof - from A_case have "\rex'. [] = [] \ Bk # rev CL @ Oc # Bk \ Suc rex = Bk # rev CL @ Oc # Bk \ Suc rex' \ CL \ [] \ last CL = Bk" (* YYYY6''' *) by simp with A_case show "inv_tm_erase_right_then_dblBk_left_erp CL CR (12, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case by simp next assume "\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc" (* YYYY2'' *) then obtain rex where A_case: "l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc" by blast then have "hd (rev CL) = Oc" by (simp add: hd_rev) with A_case have "step0 (11, [] , rev CL @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk # rev CL @ Oc # Bk \ Suc rex )" proof - from A_case have "\rex'. [] = [] \ Bk # rev CL @ Oc # Bk \ Suc rex = Bk # rev CL @ Oc # Bk \ Suc rex' \ (CL = [] \ last CL = Oc)" (* YYYY1' *) by simp with A_case show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 11\ and A_case by simp qed next assume "s = 12" with cf_cases and assms have " (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc) \ (\rex. l = [] \ r = Bk#rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk) \ (\rex. l = [] \ r = Bk#Bk#rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc))" by auto then have s12_cases: "\P. \ \rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc \ P; \rex. l = [] \ r = Bk#rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk \ P; \rex. l = [] \ r = Bk#Bk#rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc) \ P \ \ P" by blast show ?thesis proof (rule s12_cases) (* YYYY4 *) assume "\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc" then obtain rex ls1 ls2 where A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2\ tl ls1 \ [] \ last ls1 = Oc" by blast then have "ls1 \ []" by auto with A_case have major: "hd (rev ls1) = Oc" by (simp add: hd_rev) show ?thesis proof (rule noDblBk_cases) from \noDblBk CL\ show "noDblBk CL" . next from A_case show "CL = ls1 @ ls2" by auto next assume "ls2 = []" (* A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] \ last ls1 = Oc" major: "hd (rev ls1) = Oc ls1 \ [] ass: ls2 = [] l = [] r = rev ls1 @ Oc # Bk \ Suc rex where "hd (rev ls1) = Oc" \ s11 l' = [] r' = Bk#rev ls1 @ Oc # Bk \ Suc rex *) from A_case have "step0 (12, [] , Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left = (11, [], Bk#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover from A_case and major have "r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)" by (metis Nil_is_append_conv \ls1 \ []\ hd_Cons_tl hd_append2 list.simps(3) rev_is_Nil_conv) ultimately have "step0 (12, [] , rev ls1 @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" by (simp add: A_case) (* CL = ls1 @ ls2 \ tl ls1 \ [] CL = ls1 l' = [] r' = Bk# (rev CL) @ Oc # Bk \ Suc rex *) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" proof - from A_case and \ls2 = []\ have "rev ls1 = rev CL" by auto with A_case and \ls2 = []\ have "\rex'. [] = [] \ Bk# rev ls1 @ Oc # Bk \ Suc rex = Bk# rev CL @ Oc # Bk \ Suc rex' \ (CL = [] \ last CL = Oc)" (* YYYY1' *) by simp with A_case \r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)\ and \ls2 = []\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 12\ and A_case and \ls2 = []\ by simp next assume "ls2 = [Bk]" from A_case have "step0 (12, [Bk] , Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left = (11, [], Bk#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover from A_case and major have "r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)" by (metis Nil_is_append_conv \ls1 \ []\ hd_Cons_tl hd_append2 list.simps(3) rev_is_Nil_conv) ultimately have "step0 (12, [Bk] , rev ls1 @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" by (simp add: A_case) (* CL = ls1 @ ls2 \ tl ls1 \ [] CL = ls1 @ [Bk] = (ls1 @ [Bk]) = ls1' l' = [] r' = rev ls1' = rev CL l' = [] r' = rev CL @ Oc # Bk \ Suc rex *) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" proof - from A_case and \ls2 = [Bk]\ have "CL = ls1 @ [Bk]" by auto then have "\rex'. [] = [] \ Bk#rev ls1 @ Oc # Bk \ Suc rex = rev CL @ Oc # Bk \ Suc rex' \ CL \ [] \ last CL = Bk" (* YYYY2' *) by simp with A_case \r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)\ and \ls2 = [Bk]\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 12\ and A_case and \ls2 = [Bk]\ by simp next fix C3 assume minor: "ls2 = Bk # Oc # C3" (* A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] \ last ls1 = Oc" major: "hd (rev ls1) = Oc ls1 \ [] ass: ls2 = Bk # Oc # C3 l = Bk # Oc # C3 r = rev ls1 @ Oc # Bk \ Suc rex where "hd (rev ls1) = Oc" \ s11 l' = Oc # C3 r' = Bk#rev ls1 @ Oc # Bk \ Suc rex *) from A_case have "step0 (12, Bk # Oc # C3 , Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left = (11, Oc # C3, Bk#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover from A_case and major have "r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)" by (metis Nil_is_append_conv \ls1 \ []\ hd_Cons_tl hd_append2 list.simps(3) rev_is_Nil_conv) ultimately have "step0 (12, Bk # Oc # C3 , rev ls1 @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, Oc # C3, Bk# rev ls1 @ Oc # Bk \ Suc rex )" by (simp add: A_case) (* CL = ls1 @ ls2 \ tl ls1 \ [] CL = ls1 @ (Bk # Oc # C3) = ls1 @ [Bk] @ (Oc # C3) ls1' = ls1 @ [Bk] ls2' = Oc # C3 rev ls1' = Bk# (rev ls1) l' = Oc # C3 r' = rev ls1' @ Oc # Bk \ Suc rex \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ [] YYYY8 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ []) *) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, Oc # C3, Bk# rev ls1 @ Oc # Bk \ Suc rex )" proof - from A_case and \ls2 = Bk # Oc # C3\ have "CL = ls1 @ [Bk] @ (Oc # C3)" and "rev (ls1 @ [Bk]) = Bk # rev ls1" by auto with \ls1 \ []\ have "\rex' ls1' ls2'. Oc # C3 = ls2' \ Bk# rev ls1 @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ tl ls1' \ []" (* YYYY8 *) by (simp add: A_case ) with A_case \r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)\ and \ls2 = Bk # Oc # C3\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, Oc # C3, Bk# rev ls1 @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 12\ and A_case and \ls2 = Bk # Oc # C3\ by simp next fix C3 assume minor: "ls2 = Oc # C3" (* A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ [] \ last ls1 = Oc" major: "hd (rev ls1) = Oc ls1 \ [] ass: ls2 = Bk # Oc # C3 l = ls2 = Oc # C3 r = rev ls1 @ Oc # Bk \ Suc rex where "hd (rev ls1) = Oc" \ s11 l' = C3 r' = Oc#rev ls1 @ Oc # Bk \ Suc rex *) from A_case have "step0 (12, Oc # C3 , Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left = (11, C3, Oc#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover from A_case and major have "r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)" by (metis Nil_is_append_conv \ls1 \ []\ hd_Cons_tl hd_append2 list.simps(3) rev_is_Nil_conv) ultimately have "step0 (12, Oc # C3 , rev ls1 @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (11, C3, Oc# rev ls1 @ Oc # Bk \ Suc rex )" by (simp add: A_case) (* CL = ls1 @ ls2 \ tl ls1 \ [] CL = ls1 @ (Oc # C3) = ls1 @ [Oc] @ (C3) ls1' = ls1 @ [Oc] ls2' = C3 rev ls1' = Oc# (rev ls1) l' = C3 r' = rev ls1' @ Oc # Bk \ Suc rex \ CL = ls1' @ ls2' \ hd ls1' = Oc \ tl ls1' \ [] YYYY8 (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ hd ls1 = Oc \ tl ls1 \ []) *) moreover have "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, C3, Oc# rev ls1 @ Oc # Bk \ Suc rex )" proof - from A_case and \ls2 = Oc # C3\ have "CL = ls1 @ [Oc] @ (C3)" and "rev (ls1 @ [Oc]) = Oc # rev ls1" by auto with \ls1 \ []\ have "\rex' ls1' ls2'. C3 = ls2' \ Oc# rev ls1 @ Oc # Bk \ Suc rex = rev ls1' @ Oc # Bk \ Suc rex' \ CL = ls1' @ ls2' \ tl ls1' \ []" (* YYYY8 *) by (simp add: A_case ) with A_case \r = Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)\ and \ls2 = Oc # C3\ show "inv_tm_erase_right_then_dblBk_left_erp CL CR (11, C3, Oc# rev ls1 @ Oc # Bk \ Suc rex )" by force qed ultimately show ?thesis using assms and cf_cases and \s = 12\ and A_case and \ls2 = Oc # C3\ by simp qed next (* YYYY6''' *) assume "\rex. l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" then obtain rex where A_case: "l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" by blast then have "step0 (12, [] , Bk # rev CL @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (0, [], Bk # rev CL @ Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with A_case have "inv_tm_erase_right_then_dblBk_left_erp CL CR (0, [], Bk # rev CL @ Oc # Bk \ Suc rex)" by auto ultimately show ?thesis using assms and cf_cases and \s = 12\ and A_case by simp next (* YYYY9 *) assume "\rex. l = [] \ r = Bk # Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" then obtain rex where A_case: "l = [] \ r = Bk # Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" by blast then have "step0 (12, [] , Bk # Bk # rev CL @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left = (0, [], Bk# Bk # rev CL @ Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with A_case have "inv_tm_erase_right_then_dblBk_left_erp CL CR (0, [], Bk# Bk # rev CL @ Oc # Bk \ Suc rex)" by auto ultimately show ?thesis using assms and cf_cases and \s = 12\ and A_case by simp qed next assume "s = 0" with cf_cases and assms have "(\rex. l = [] \ r = [Bk, Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex \ (CL = [] \ last CL = Oc) ) \ (\rex. l = [] \ r = [Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex \ CL \ [] \ last CL = Bk)" by auto then have s0_cases: "\P. \ \rex. l = [] \ r = [Bk, Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex \ (CL = [] \ last CL = Oc) \ P; \rex. l = [] \ r = [Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex \ CL \ [] \ last CL = Bk \ P \ \ P" by blast show ?thesis proof (rule s0_cases) assume "\rex. l = [] \ r = [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex \ (CL = [] \ last CL = Oc)" then obtain rex where A_case: "l = [] \ r = [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex \ (CL = [] \ last CL = Oc)" by blast then have "step0 (0, [] , [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex) tm_erase_right_then_dblBk_left = (0, [], Bk# Bk # rev CL @ Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with A_case have "inv_tm_erase_right_then_dblBk_left_erp CL CR (0, [], Bk# Bk # rev CL @ Oc # Bk \ Suc rex)" by auto ultimately show ?thesis using assms and cf_cases and \s = 0\ and A_case by simp next assume "\rex. l = [] \ r = [Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex \ CL \ [] \ last CL = Bk" then obtain rex where A_case: "l = [] \ r = [Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex \ CL \ [] \ last CL = Bk" by blast then have "step0 (0, [] , [Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex) tm_erase_right_then_dblBk_left = (0, [], Bk # rev CL @ Oc # Bk \ Suc rex)" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) moreover with A_case have "inv_tm_erase_right_then_dblBk_left_erp CL CR (0, [], Bk # rev CL @ Oc # Bk \ Suc rex)" by auto ultimately show ?thesis using assms and cf_cases and \s = 0\ and A_case by simp qed qed qed (* -------------- steps - lemma for the invariants of the ERASE path of tm_erase_right_then_dblBk_left ------------------ *) lemma inv_tm_erase_right_then_dblBk_left_erp_steps: assumes "inv_tm_erase_right_then_dblBk_left_erp CL CR cf" and "noDblBk CL" and "noDblBk CR" shows "inv_tm_erase_right_then_dblBk_left_erp CL CR (steps0 cf tm_erase_right_then_dblBk_left stp)" proof (induct stp) case 0 with assms show ?case by (auto simp add: inv_tm_erase_right_then_dblBk_left_erp_step step.simps steps.simps) next case (Suc stp) with assms show ?case using inv_tm_erase_right_then_dblBk_left_erp_step step_red by auto qed (* -------------- Partial correctness for the ERASE path of tm_erase_right_then_dblBk_left ------------------ *) lemma tm_erase_right_then_dblBk_left_erp_partial_correctness_CL_is_Nil: assumes "\stp. is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" and "noDblBk CL" and "noDblBk CR" and "CL = []" shows "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" proof (rule Hoare_consequence) show "( \tap. tap = ([Bk,Oc] @ CL, CR) ) \ ( \tap. tap = ([Bk,Oc] @ CL, CR) )" by auto next from assms show "inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR \ ( \tap. \rex. tap = ([], [Bk,Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex) )" by (simp add: assert_imp_def tape_of_list_def tape_of_nat_def) next show " \\tap. tap = ([Bk, Oc] @ CL, CR)\ tm_erase_right_then_dblBk_left \inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR\" proof (rule Hoare_haltI) fix l::"cell list" fix r:: "cell list" assume major: "(l, r) = ([Bk, Oc] @ CL, CR)" show "\n. is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left n) \ inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR holds_for steps0 (1, l, r) tm_erase_right_then_dblBk_left n" proof - from major and assms have "\stp. is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" by blast then obtain stp where w_stp: "is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" by blast moreover have "inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR holds_for steps0 (1, l, r) tm_erase_right_then_dblBk_left stp" proof - have "inv_tm_erase_right_then_dblBk_left_erp CL CR (1, l, r)" by (simp add: major tape_of_list_def tape_of_nat_def) with assms have "inv_tm_erase_right_then_dblBk_left_erp CL CR (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" using inv_tm_erase_right_then_dblBk_left_erp_steps by auto then show ?thesis - by (smt holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_erp.simps is_final_eq w_stp) + by (smt (verit) holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_erp.simps is_final_eq w_stp) qed ultimately show ?thesis by auto qed qed qed lemma tm_erase_right_then_dblBk_left_erp_partial_correctness_CL_ew_Bk: assumes "\stp. is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" and "noDblBk CL" and "noDblBk CR" and "CL \ []" and "last CL = Bk" shows "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" proof (rule Hoare_consequence) show "( \tap. tap = ([Bk,Oc] @ CL, CR) ) \ ( \tap. tap = ([Bk,Oc] @ CL, CR) )" by auto next from assms show "inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR \ ( \tap. \rex. tap = ([], [Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex) )" by (simp add: assert_imp_def tape_of_list_def tape_of_nat_def) next show " \\tap. tap = ([Bk, Oc] @ CL, CR)\ tm_erase_right_then_dblBk_left \inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR\" proof (rule Hoare_haltI) fix l::"cell list" fix r:: "cell list" assume major: "(l, r) = ([Bk, Oc] @ CL, CR)" show "\n. is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left n) \ inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR holds_for steps0 (1, l, r) tm_erase_right_then_dblBk_left n" proof - from major and assms have "\stp. is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" by blast then obtain stp where w_stp: "is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" by blast moreover have "inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR holds_for steps0 (1, l, r) tm_erase_right_then_dblBk_left stp" proof - have "inv_tm_erase_right_then_dblBk_left_erp CL CR (1, l, r)" by (simp add: major tape_of_list_def tape_of_nat_def) with assms have "inv_tm_erase_right_then_dblBk_left_erp CL CR (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" using inv_tm_erase_right_then_dblBk_left_erp_steps by auto then show ?thesis - by (smt holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_erp.simps is_final_eq w_stp) + by (smt (verit) holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_erp.simps is_final_eq w_stp) qed ultimately show ?thesis by auto qed qed qed lemma tm_erase_right_then_dblBk_left_erp_partial_correctness_CL_ew_Oc: assumes "\stp. is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" and "noDblBk CL" and "noDblBk CR" and "CL \ []" and "last CL = Oc" shows "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk, Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" proof (rule Hoare_consequence) show "( \tap. tap = ([Bk,Oc] @ CL, CR) ) \ ( \tap. tap = ([Bk,Oc] @ CL, CR) )" by auto next from assms show "inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR \ ( \tap. \rex. tap = ([], [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex) )" by (simp add: assert_imp_def tape_of_list_def tape_of_nat_def) next show " \\tap. tap = ([Bk, Oc] @ CL, CR)\ tm_erase_right_then_dblBk_left \inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR\" proof (rule Hoare_haltI) fix l::"cell list" fix r:: "cell list" assume major: "(l, r) = ([Bk, Oc] @ CL, CR)" show "\n. is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left n) \ inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR holds_for steps0 (1, l, r) tm_erase_right_then_dblBk_left n" proof - from major and assms have "\stp. is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" by blast then obtain stp where w_stp: "is_final (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" by blast moreover have "inv_tm_erase_right_then_dblBk_left_erp_s0 CL CR holds_for steps0 (1, l, r) tm_erase_right_then_dblBk_left stp" proof - have "inv_tm_erase_right_then_dblBk_left_erp CL CR (1, l, r)" by (simp add: major tape_of_list_def tape_of_nat_def) with assms have "inv_tm_erase_right_then_dblBk_left_erp CL CR (steps0 (1, l, r) tm_erase_right_then_dblBk_left stp)" using inv_tm_erase_right_then_dblBk_left_erp_steps by auto then show ?thesis - by (smt holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_erp.simps is_final_eq w_stp) + by (smt (verit) holds_for.elims(3) inv_tm_erase_right_then_dblBk_left_erp.simps is_final_eq w_stp) qed ultimately show ?thesis by auto qed qed qed (* -------------- Termination for the ERASE path of tm_erase_right_then_dblBk_left ------------------ *) (* Lexicographic orders (See List.measures) quote: "These are useful for termination proofs" lemma in_measures[simp]: "(x, y) \ measures [] = False" "(x, y) \ measures (f # fs) = (f x < f y \ (f x = f y \ (x, y) \ measures fs))" *) (* Assemble a lexicographic measure function for the ERASE path *) definition measure_tm_erase_right_then_dblBk_left_erp :: "(config \ config) set" where "measure_tm_erase_right_then_dblBk_left_erp = measures [ \(s, l, r). ( if s = 0 then 0 else if s < 6 then 13 - s else 1), \(s, l, r). ( if s = 6 then if r = [] \ (hd r) = Bk then 1 else 2 else 0 ), \(s, l, r). ( if 7 \ s \ s \ 9 then 2+ length r else 1), \(s, l, r). ( if 7 \ s \ s \ 9 then if r = [] \ hd r = Bk then 2 else 3 else 1), \(s, l, r).( if 7 \ s \ s \ 10 then 13 - s else 1), \(s, l, r). ( if 10 \ s then 2+ length l else 1), \(s, l, r). ( if 11 \ s then if hd r = Oc then 3 else 2 else 1), \(s, l, r).( if 11 \ s then 13 - s else 1) ]" lemma wf_measure_tm_erase_right_then_dblBk_left_erp: "wf measure_tm_erase_right_then_dblBk_left_erp" unfolding measure_tm_erase_right_then_dblBk_left_erp_def by (auto) lemma measure_tm_erase_right_then_dblBk_left_erp_induct [case_names Step]: "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_tm_erase_right_then_dblBk_left_erp\ \ \n. P (f n)" using wf_measure_tm_erase_right_then_dblBk_left_erp by (metis wf_iff_no_infinite_down_chain) lemma spike_erp_cases: "CL \ [] \ last CL = Bk \ CL \ [] \ last CL = Oc \ CL = []" using cell.exhaust by blast lemma tm_erase_right_then_dblBk_left_erp_halts: assumes "noDblBk CL" and "noDblBk CR" shows "\stp. is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" proof (induct rule: measure_tm_erase_right_then_dblBk_left_erp_induct) case (Step stp) then have not_final: "\ is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" . have INV: "inv_tm_erase_right_then_dblBk_left_erp CL CR (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" proof (rule_tac inv_tm_erase_right_then_dblBk_left_erp_steps) show "inv_tm_erase_right_then_dblBk_left_erp CL CR (1, [Bk,Oc] @ CL, CR)" by (simp add: tape_of_list_def tape_of_nat_def ) next from assms show "noDblBk CL" by auto next from assms show "noDblBk CR" by auto qed have SUC_STEP_RED: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp) tm_erase_right_then_dblBk_left" by (rule step_red) show "( steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp), steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp ) \ measure_tm_erase_right_then_dblBk_left_erp" proof (cases "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp") case (fields s l r) then have cf_at_stp: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (s, l, r)" . show ?thesis proof (rule tm_erase_right_then_dblBk_left_erp_cases) from INV and cf_at_stp show "inv_tm_erase_right_then_dblBk_left_erp CL CR (s, l, r)" by auto next assume "s=0" (* not possible *) with cf_at_stp not_final show ?thesis by auto next assume "s=1" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (1, l, r)" by auto with cf_at_stp and \s=1\ and INV have unpacked_INV: "(l = [Bk,Oc] @ CL \ r = CR)" by auto (* compute the next state *) show ?thesis proof (cases CR) case Nil then have minor: "CR = []" . with unpacked_INV cf_at_stp and \s=1\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (2,Oc#CL, Bk#CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (2,Oc#CL, Bk#CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case (Cons a rs) then have major: "CR = a # rs" . then show ?thesis proof (cases a) case Bk with major have minor: "CR = Bk#rs" by auto with unpacked_INV cf_at_stp and \s=1\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (2,Oc#CL, Bk#CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (2,Oc#CL, Bk#CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case Oc with major have minor: "CR = Oc#rs" by auto with unpacked_INV cf_at_stp and \s=1\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (2,Oc#CL, Bk#CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (2,Oc#CL, Bk#CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed qed next assume "s=2" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (2, l, r)" by auto with cf_at_stp and \s=2\ and INV have unpacked_INV: "(l = [Oc] @ CL \ r = Bk#CR)" by auto (* compute the next state *) then have minor: "r = Bk#CR" by auto with unpacked_INV cf_at_stp and \s=2\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (2, [Oc] @ CL, Bk#CR) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (3,CL, Oc#Bk#CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (3,CL, Oc#Bk#CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "s=3" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (3, l, r)" by auto with cf_at_stp and \s=3\ and INV have unpacked_INV: "(l = CL \ r = Oc#Bk#CR)" by auto (* compute the next state *) then have minor: "r = Oc#Bk#CR" by auto with unpacked_INV cf_at_stp and \s=3\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (3, CL, Oc#Bk#CR) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (5,[Oc] @ CL, Bk#CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (5,[Oc] @ CL, Bk#CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "s=5" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (5, l, r)" by auto with cf_at_stp and \s=5\ and INV have unpacked_INV: "(l = [Oc] @ CL \ r = Bk#CR)" by auto (* compute the next state *) then have minor: "r = Bk#CR" by auto with unpacked_INV cf_at_stp and \s=5\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (5, [Oc] @ CL, Bk#CR) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (6, [Bk,Oc] @ CL, CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (6, [Bk,Oc] @ CL, CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "s=6" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (6, l, r)" by auto with cf_at_stp and \s=6\ and INV have unpacked_INV: "(l = [Bk,Oc] @ CL \ ( (CR = [] \ r = CR) \ (CR \ [] \ (r = CR \ r = Bk # tl CR)) ))" by auto then have unpacked_INV': "l = [Bk,Oc] @ CL \ CR = [] \ r = CR \ l = [Bk,Oc] @ CL \ CR \ [] \ r = Oc # tl CR \ l = [Bk,Oc] @ CL \ CR \ [] \ r = Bk # tl CR" by (metis (full_types) cell.exhaust list.sel(3) neq_Nil_conv) then show ?thesis proof assume minor: "l = [Bk, Oc] @ CL \ CR = [] \ r = CR" with unpacked_INV cf_at_stp and \s=6\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (6, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (7,Bk#[Bk, Oc] @ CL, CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (7,Bk#[Bk, Oc] @ CL, CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "l = [Bk, Oc] @ CL \ CR \ [] \ r = Oc # tl CR \ l = [Bk, Oc] @ CL \ CR \ [] \ r = Bk # tl CR" then show ?thesis proof assume minor: "l = [Bk, Oc] @ CL \ CR \ [] \ r = Bk # tl CR" with unpacked_INV cf_at_stp and \s=6\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (6, [Bk, Oc] @ CL, Bk # tl CR) tm_erase_right_then_dblBk_left" by auto also with minor have "... = (7,Bk#[Bk, Oc] @ CL, tl CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (7,Bk#[Bk, Oc] @ CL, tl CR)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume minor: "l = [Bk, Oc] @ CL \ CR \ [] \ r = Oc # tl CR" with unpacked_INV cf_at_stp and \s=6\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (6, [Bk, Oc] @ CL, Oc # tl CR) tm_erase_right_then_dblBk_left" by auto also with minor have "... = (6, [Bk, Oc] @ CL, Bk # tl CR)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (6, [Bk, Oc] @ CL, Bk # tl CR)" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed qed next assume "s=7" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (7, l, r)" by auto with cf_at_stp and \s=7\ and INV have "(\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs. CR = rs @ r)" by auto then obtain lex rs where unpacked_INV: "l = Bk \ Suc lex @ [Bk,Oc] @ CL \ CR = rs @ r" by blast (* compute the next state *) show ?thesis proof (cases r) case Nil then have minor: "r = []" . with unpacked_INV cf_at_stp and \s=7\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (7, Bk \ Suc lex @ [Bk,Oc] @ CL, r) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, r)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, r)" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case (Cons a rs') then have major: "r = a # rs'" . then show ?thesis proof (cases a) case Bk with major have minor: "r = Bk#rs'" by auto with unpacked_INV cf_at_stp and \s=7\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (7, Bk \ Suc lex @ [Bk,Oc] @ CL, r) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs')" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (9, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs')" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case Oc with major have minor: "r = Oc#rs'" by auto with unpacked_INV cf_at_stp and \s=7\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (7, Bk \ Suc lex @ [Bk,Oc] @ CL, Oc#rs') tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed qed next assume "s=8" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (8, l, r)" by auto with cf_at_stp and \s=8\ and INV have "(\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs1 rs2. CR = rs1 @ [Oc] @ rs2 \ r = Bk#rs2)" by auto then obtain lex rs1 rs2 where unpacked_INV: "l = Bk \ Suc lex @ [Bk,Oc] @ CL \ CR = rs1 @ [Oc] @ rs2 \ r = Bk#rs2 " by blast (* compute the next state *) with cf_at_stp and \s=8\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs2) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (7, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs2)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (7, Bk \ Suc (Suc lex) @ [Bk,Oc] @ CL, rs2)" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "s=9" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (9, l, r)" by auto with cf_at_stp and \s=9\ and INV have "(\lex. l = Bk \ Suc lex @ [Bk,Oc] @ CL) \ (\rs. CR = rs @ [Bk] @ r \ CR = rs \ r = [])" by auto then obtain lex rs where "l = Bk \ Suc lex @ [Bk,Oc] @ CL \ (CR = rs @ [Bk] @ r \ CR = rs \ r = [])" by blast then have unpacked_INV: "l = Bk \ Suc lex @ [Bk,Oc] @ CL \ CR = rs @ [Bk] @ r \ l = Bk \ Suc lex @ [Bk,Oc] @ CL \ CR = rs \ r = []" by auto then show ?thesis proof (* compute the next state *) assume major: "l = Bk \ Suc lex @ [Bk, Oc] @ CL \ CR = rs @ [Bk] @ r" show ?thesis proof (cases r) case Nil then have minor: "r = []" . with unpacked_INV cf_at_stp and \s=9\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (9, Bk \ Suc lex @ [Bk,Oc] @ CL, r) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (10, Bk \ lex @ [Bk,Oc] @ CL, Bk#r)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (10, Bk \ lex @ [Bk,Oc] @ CL, Bk#r)" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case (Cons a rs') then have major2: "r = a # rs'" . then show ?thesis proof (cases a) case Bk with major2 have minor: "r = Bk#rs'" by auto with unpacked_INV cf_at_stp and \s=9\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (9, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs') tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (10, Bk \ lex @ [Bk,Oc] @ CL, Bk#Bk#rs')" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (10, Bk \ lex @ [Bk,Oc] @ CL, Bk#Bk#rs')" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case Oc with major2 have minor: "r = Oc#rs'" by auto with unpacked_INV cf_at_stp and \s=9\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (9, Bk \ Suc lex @ [Bk,Oc] @ CL, Oc#rs') tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (8, Bk \ Suc lex @ [Bk,Oc] @ CL, Bk#rs')" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed qed next (* compute the next state *) assume major: "l = Bk \ Suc lex @ [Bk, Oc] @ CL \ CR = rs \ r = []" with unpacked_INV cf_at_stp and \s=9\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (9, Bk \ Suc lex @ [Bk,Oc] @ CL, []) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (10, Bk \ lex @ [Bk,Oc] @ CL, [Bk])" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed next assume "s=10" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (10, l, r)" by auto with cf_at_stp and \s=10\ and INV have "(\lex rex. l = Bk \ lex @ [Bk,Oc] @ CL \ r = Bk \ Suc rex) \ (\rex. l = [Oc] @ CL \ r = Bk \ Suc rex) \ (\rex. l = CL \ r = Oc # Bk \ Suc rex)" by auto then have s10_cases: "\P. \ \lex rex. l = Bk \ lex @ [Bk,Oc] @ CL \ r = Bk \ Suc rex \ P; \rex. l = [Oc] @ CL \ r = Bk \ Suc rex \ P; \rex. l = CL \ r = Oc # Bk \ Suc rex \ P \ \ P" by blast show ?thesis proof (rule s10_cases) assume "\lex rex. l = Bk \ lex @ [Bk, Oc] @ CL \ r = Bk \ Suc rex" then obtain lex rex where unpacked_INV: "l = Bk \ lex @ [Bk, Oc] @ CL \ r = Bk \ Suc rex" by blast with unpacked_INV cf_at_stp and \s=10\ and SUC_STEP_RED have todo_step: "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (10, Bk \ lex @ [Bk, Oc] @ CL, Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto show ?thesis proof (cases lex) case 0 then have "lex = 0" . (* compute the next state *) then have "step0 (10, Bk \ lex @ [Bk, Oc] @ CL, Bk \ Suc rex) tm_erase_right_then_dblBk_left = (10, [Oc] @ CL, Bk \ Suc (Suc rex))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) with todo_step have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (10, [Oc] @ CL, Bk \ Suc (Suc rex))" by auto (* establish measure *) with \lex = 0\ and cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case (Suc nat) then have "lex = Suc nat" . (* compute the next state *) then have "step0 (10, Bk \ lex @ [Bk, Oc] @ CL, Bk \ Suc rex) tm_erase_right_then_dblBk_left = (10, Bk \ nat @ [Bk, Oc] @ CL, Bk \ Suc (Suc rex))" by (simp add: tm_erase_right_then_dblBk_left_def step.simps steps.simps numeral_eqs_upto_12) with todo_step have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (10, Bk \ nat @ [Bk, Oc] @ CL, Bk \ Suc (Suc rex))" by auto (* establish measure *) with \lex = Suc nat\ cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed next assume "\rex. l = [Oc] @ CL \ r = Bk \ Suc rex" then obtain rex where unpacked_INV: "l = [Oc] @ CL \ r = Bk \ Suc rex" by blast (* compute the next state *) with unpacked_INV cf_at_stp and \s=10\ and SUC_STEP_RED have todo_step: "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (10, [Oc] @ CL, Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (10, CL, Oc# Bk \ (Suc rex))" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (10, CL, Oc# Bk \ (Suc rex))" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "\rex. l = CL \ r = Oc # Bk \ Suc rex" then obtain rex where unpacked_INV: "l = CL \ r = Oc # Bk \ Suc rex" by blast show ?thesis proof (cases CL) (* here, we start decomposing CL in the loop 'move to left until Oc *) case Nil then have minor: "CL = []" . with unpacked_INV cf_at_stp and \s=10\ and SUC_STEP_RED (* compute the next state *) have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (10, [], Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (11, [], Bk#Oc# Bk \ (Suc rex))" (* YYYY1' *) by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, [], Bk#Oc# Bk \ (Suc rex))" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case (Cons a rs') then have major2: "CL = a # rs'" . then show ?thesis proof (cases a) case Bk with major2 have minor: "CL = Bk#rs'" by auto with unpacked_INV cf_at_stp and \s=10\ and SUC_STEP_RED (* compute the next state *) have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (10, Bk#rs', Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (11, rs', Bk# Oc # Bk \ Suc rex)" (* YYYY2',YYYY3 *) by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, rs', Bk# Oc # Bk \ Suc rex)" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next case Oc with major2 have minor: "CL = Oc#rs'" by auto with unpacked_INV cf_at_stp and \s=10\ and SUC_STEP_RED (* compute the next state *) have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (10, Oc#rs', Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also with minor and unpacked_INV have "... = (11, rs', Oc# Oc # Bk \ Suc rex)" (* YYYY2'', YYYY5, YYYY6, YYYY7 *) by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, rs', Oc# Oc # Bk \ Suc rex)" by auto (* establish measure *) with cf_at_current and minor show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed qed qed next assume "s=11" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (11, l, r)" by auto with cf_at_stp and \s=11\ and INV have "(\rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)) \ (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk ) \ (\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc ) \ (\rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]) \ (\rex ls1 ls2. l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc] ) \ (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk] ) \ (\rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc] ) \ (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [])" by auto then have s11_cases: "\P. \ \rex. l = [] \ r = Bk# rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc) \ P; \rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk \ P; \rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc \ P; \rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk] \ P; \rex ls1 ls2. l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc] \ P; \rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Bk] \ P; \rex ls1 ls2. l = Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc#ls2 \ ls1 = [Oc] \ P; \rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ P \ \ P" by blast show ?thesis proof (rule s11_cases) assume "\rex ls1 ls2. l = Bk # Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk # Oc # ls2 \ ls1 = [Oc]" (* YYYY6 *) then obtain rex ls1 ls2 where unpacked_INV: "l = Bk#Oc#ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Bk#Oc#ls2 \ ls1 = [Oc]" by blast from unpacked_INV cf_at_stp and \s=11\ and SUC_STEP_RED have todo_step: "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, Bk#Oc#ls2, r) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (11, Oc # ls2, Bk # r)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, Oc # ls2, Bk # r)" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "\rex ls1 ls2. l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Oc]" (* YYYY7 *) then obtain rex ls1 ls2 where unpacked_INV: "l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Oc]" by blast from unpacked_INV cf_at_stp and \s=11\ and SUC_STEP_RED have todo_step: "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, l, r) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (11, ls2, Oc#r)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, ls2, Oc#r)" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "\rex. l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]" (* YYYY5 *) then obtain rex where unpacked_INV: "l = [Bk] \ r = rev [Oc] @ Oc # Bk \ Suc rex \ CL = [Oc, Bk]" by blast (* compute the next state *) from unpacked_INV cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, l, r) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (11, [], Bk#r)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, [], Bk#r)" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ []" (* YYYY8 *) then obtain rex ls1 ls2 where A_case: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ []" by blast (* we have to decompose both ls1 and ls2 in order to predict the next step *) then have "\z b bs. ls1 = z#bs@[b]" (* the symbol z does not matter *) by (metis Nil_tl list.exhaust_sel rev_exhaust) then have "\z bs. ls1 = z#bs@[Bk] \ ls1 = z#bs@[Oc]" using cell.exhaust by blast then obtain z bs where w_z_bs: "ls1 = z#bs@[Bk] \ ls1 = z#bs@[Oc]" by blast then show ?thesis proof assume major1: "ls1 = z # bs @ [Bk]" then have major2: "rev ls1 = Bk#(rev bs)@[z]" by auto (* in this case all transitions will be s11 \ s12 *) show ?thesis proof (rule noDblBk_cases) from \noDblBk CL\ show "noDblBk CL" . next from A_case show "CL = ls1 @ ls2" by auto next assume minor: "ls2 = []" (* compute the next state *) with A_case major2 cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, [], Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (12, [], Bk#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (12, [], Bk#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "ls2 = [Bk]" with A_case \rev ls1 = Bk#(rev bs)@[z]\ and \ls2 = [Bk]\ have "ls1 = z#bs@[Bk]" and "CL = z#bs@[Bk]@[Bk]" by auto with \noDblBk CL\ have False by (metis A_case \ls2 = [Bk]\ append_Cons hasDblBk_L5 major2) then show ?thesis by auto next fix C3 assume minor: "ls2 = Bk # Oc # C3" with A_case and major2 have "CL = z # bs @ [Bk] @ Bk # Oc # C3" by auto with \noDblBk CL\ have False by (metis append.left_neutral append_Cons append_assoc hasDblBk_L1 major1 minor) then show ?thesis by auto next fix C3 assume minor: "ls2 = Oc # C3" (* compute the next state *) with A_case major2 cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, Oc # C3 , Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (12, C3, Oc#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (12, C3, Oc#Bk#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with A_case minor cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed next assume major1: "ls1 = z # bs @ [Oc]" then have major2: "rev ls1 = Oc#(rev bs)@[z]" by auto (* in this case all transitions will be s11 \ s11 *) show ?thesis proof (rule noDblBk_cases) from \noDblBk CL\ show "noDblBk CL" . next from A_case show "CL = ls1 @ ls2" by auto next assume minor: "ls2 = []" (* compute the next state *) with A_case major2 cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, [] , Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex)" by auto (* establish measure *) with A_case minor major2 cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume minor: "ls2 = [Bk]" (* compute the next state *) with A_case major2 cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, [Bk] , Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, [], Bk#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with A_case minor major2 cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next fix C3 assume minor: "ls2 = Bk # Oc # C3" (* compute the next state *) with A_case major2 cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, Bk # Oc # C3 , Oc # rev bs @ [z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (11, Oc # C3, Bk#Oc # rev bs @ [z] @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, Oc # C3, Bk#Oc # rev bs @ [z] @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with A_case minor major2 cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next fix C3 assume minor: "ls2 = Oc # C3" (* compute the next state *) with A_case major2 cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, Oc # C3 , Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (11, C3, Oc#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, C3, Oc#Oc#(rev bs)@[z] @ Oc # Bk \ Suc rex)" by auto (* establish measure *) with A_case minor major2 cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed qed next assume "\rex ls1 ls2. l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Bk]" (* YYYY3 *) then obtain rex ls1 ls2 where unpacked_INV: "l = Oc # ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ Oc # ls2 \ ls1 = [Bk]" by blast (* compute the next state *) from unpacked_INV cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, l, r) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (12, ls2, Oc#[Bk] @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (12, ls2, Oc#[Bk] @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "\rex. l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" (* YYYY1' *) then obtain rex where unpacked_INV: "l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" by blast (* compute the next state *) from unpacked_INV cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, l, r) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV have "... = (12, [], Bk#Bk # rev CL @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (12, [], Bk#Bk # rev CL @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" (* YYYY2' *) then obtain rex where unpacked_INV: "l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" by blast then have "hd (rev CL) = Bk" by (simp add: hd_rev) (* compute the next state *) from unpacked_INV cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, l, r) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV and \hd (rev CL) = Bk\ have "... = (12, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (12, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with cf_at_current and unpacked_INV show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume "\rex. l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc" (* YYYY2'' *) then obtain rex where unpacked_INV: "l = [] \ r = rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Oc" by blast then have "hd (rev CL) = Oc" by (simp add: hd_rev) (* compute the next state *) from unpacked_INV cf_at_stp and \s=11\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (11, l, r) tm_erase_right_then_dblBk_left" by auto also with unpacked_INV and \hd (rev CL) = Oc\ have "... = (11, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, [], Bk # rev CL @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with cf_at_current and unpacked_INV and \hd (rev CL) = Oc\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed next assume "s=12" (* get the invariant of the state *) with cf_at_stp have cf_at_current: "steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp = (12, l, r)" by auto with cf_at_stp and \s=12\ and INV have " (\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc) \ (\rex. l = [] \ r = Bk#rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk) \ (\rex. l = [] \ r = Bk#Bk#rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc))" by auto then have s12_cases: "\P. \ \rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc \ P; \rex. l = [] \ r = Bk#rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk \ P; \rex. l = [] \ r = Bk#Bk#rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc) \ P \ \ P" by blast show ?thesis proof (rule s12_cases) (* YYYY4 *) assume "\rex ls1 ls2. l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2 \ tl ls1 \ [] \ last ls1 = Oc" then obtain rex ls1 ls2 where unpacked_INV: "l = ls2 \ r = rev ls1 @ Oc # Bk \ Suc rex \ CL = ls1 @ ls2\ tl ls1 \ [] \ last ls1 = Oc" by blast then have "ls1 \ []" by auto with unpacked_INV have major: "hd (rev ls1) = Oc" by (simp add: hd_rev) with unpacked_INV and major have minor2: "r = Oc#tl ((rev ls1) @ Oc # Bk \ Suc rex)" by (metis Nil_is_append_conv \ls1 \ []\ hd_Cons_tl hd_append2 list.simps(3) rev_is_Nil_conv) show ?thesis proof (rule noDblBk_cases) from \noDblBk CL\ show "noDblBk CL" . next from unpacked_INV show "CL = ls1 @ ls2" by auto next assume minor: "ls2 = []" (* compute the next state *) with unpacked_INV minor minor2 major cf_at_stp and \s=12\ and \ls1 \ []\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (12, [] , Oc#tl ((rev ls1) @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left" by auto also have "... = (11, [], Bk#Oc#tl ((rev ls1) @ Oc # Bk \ Suc rex))" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) also with unpacked_INV and minor2 have "... = (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" by auto finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, [], Bk# rev ls1 @ Oc # Bk \ Suc rex )" by auto (* establish measure *) with unpacked_INV minor major minor2 cf_at_current \ls1 \ []\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next assume minor: "ls2 = [Bk]" (* compute the next state *) with unpacked_INV minor minor2 major cf_at_stp and \s=12\ and \ls1 \ []\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (12, [Bk] , Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left" by auto also have "... = (11, [], Bk#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, [], Bk#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by auto (* establish measure *) with unpacked_INV minor major minor2 cf_at_current \ls1 \ []\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next fix C3 assume minor: "ls2 = Bk # Oc # C3" (* compute the next state *) with unpacked_INV minor minor2 major cf_at_stp and \s=12\ and \ls1 \ []\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (12, Bk # Oc # C3 , Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left" by auto also have "... = (11, Oc # C3, Bk#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, Oc # C3, Bk#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by auto (* establish measure *) with unpacked_INV minor major minor2 cf_at_current \ls1 \ []\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next fix C3 assume minor: "ls2 = Oc # C3" (* compute the next state *) with unpacked_INV minor minor2 major cf_at_stp and \s=12\ and \ls1 \ []\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (12, Oc # C3 , Oc#tl (rev ls1 @ Oc # Bk \ Suc rex)) tm_erase_right_then_dblBk_left" by auto also have "... = (11, C3, Oc#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (11, C3, Oc#Oc#tl (rev ls1 @ Oc # Bk \ Suc rex ))" by auto (* establish measure *) with unpacked_INV minor major minor2 cf_at_current \ls1 \ []\ show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed next (* YYYY6''' *) assume "\rex. l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" then obtain rex where unpacked_INV: "l = [] \ r = Bk # rev CL @ Oc # Bk \ Suc rex \ CL \ [] \ last CL = Bk" by blast (* compute the next state *) with cf_at_stp and \s=12\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (12, [] , Bk # rev CL @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (0, [], Bk # rev CL @ Oc # Bk \ Suc rex)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (0, [], Bk # rev CL @ Oc # Bk \ Suc rex)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) next (* YYYY9 *) assume "\rex. l = [] \ r = Bk # Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" then obtain rex where unpacked_INV: "l = [] \ r = Bk # Bk # rev CL @ Oc # Bk \ Suc rex \ (CL = [] \ last CL = Oc)" by blast (* compute the next state *) with cf_at_stp and \s=12\ and SUC_STEP_RED have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = step0 (12, [] , Bk # Bk # rev CL @ Oc # Bk \ Suc rex) tm_erase_right_then_dblBk_left" by auto also have "... = (0, [], Bk# Bk # rev CL @ Oc # Bk \ Suc rex)" by (auto simp add: tm_erase_right_then_dblBk_left_def numeral_eqs_upto_12 step.simps steps.simps) finally have "steps0 (1, [Bk, Oc] @ CL, CR) tm_erase_right_then_dblBk_left (Suc stp) = (0, [], Bk# Bk # rev CL @ Oc # Bk \ Suc rex)" by auto (* establish measure *) with cf_at_current show ?thesis by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def) qed qed qed qed (* -------------- Total correctness for the ERASE path of tm_erase_right_then_dblBk_left ------------------ *) lemma tm_erase_right_then_dblBk_left_erp_total_correctness_CL_is_Nil: assumes "noDblBk CL" and "noDblBk CR" and "CL = []" shows "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" proof (rule tm_erase_right_then_dblBk_left_erp_partial_correctness_CL_is_Nil) from assms show "\stp. is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" using tm_erase_right_then_dblBk_left_erp_halts by auto next from assms show "noDblBk CL" by auto next from assms show "noDblBk CR" by auto next from assms show "CL = []" by auto qed lemma tm_erase_right_then_dblBk_left_correctness_CL_ew_Bk: assumes "noDblBk CL" and "noDblBk CR" and "CL \ []" and "last CL = Bk" shows "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" proof (rule tm_erase_right_then_dblBk_left_erp_partial_correctness_CL_ew_Bk) from assms show "\stp. is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" using tm_erase_right_then_dblBk_left_erp_halts by auto next from assms show "noDblBk CL" by auto next from assms show "noDblBk CR" by auto next from assms show "CL \ []" by auto next from assms show "last CL = Bk" by auto qed lemma tm_erase_right_then_dblBk_left_erp_total_correctness_CL_ew_Oc: assumes "noDblBk CL" and "noDblBk CR" and "CL \ []" and "last CL = Oc" shows "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk, Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" proof (rule tm_erase_right_then_dblBk_left_erp_partial_correctness_CL_ew_Oc) from assms show "\stp. is_final (steps0 (1, [Bk,Oc] @ CL, CR) tm_erase_right_then_dblBk_left stp)" using tm_erase_right_then_dblBk_left_erp_halts by auto next from assms show "noDblBk CL" by auto next from assms show "noDblBk CR" by auto next from assms show "CL \ []" by auto next from assms show "last CL = Oc" by auto qed (* --- prove some helper theorems to convert results to lists of numeral --- *) (*--------------------------------------------------------------------------------- *) (* simple case, where we tried to check for exactly one argument and failed *) (*--------------------------------------------------------------------------------- *) lemma tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_eq_Nil_n_eq_1_last_eq_0: assumes "(nl::nat list) \ []" and "n=1" and "n \ length nl" and "last (take n nl) = 0" shows "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL = [] \ CR = () \ noDblBk CR" proof - have "rev() = " by (rule rev_numeral_list) also with assms have "... = " by (metis append_butlast_last_id take_eq_Nil zero_neq_one) also have "... = < (rev[(last (take n nl))]) @ (rev ( butlast (take n nl)))>" by simp also with assms have "... = < (rev [0]) @ (rev ( butlast (take n nl)))>" by auto finally have major: "rev() = <(rev [0]) @ (rev ( butlast (take n nl)))>" by auto with assms have "butlast (take n nl) = []" by (simp add: butlast_take) then have "<(rev [0::nat]) @ (rev ( butlast (take n nl)))> = <(rev [0::nat]) @ (rev [])>" by auto also have "... = <(rev [0::nat])>" by auto also have "... = <[0::nat]>" by auto also have "... = [Oc]" by (simp add: tape_of_list_def tape_of_nat_def ) finally have "<(rev [0::nat]) @ (rev ( butlast (take n nl)))> = [Oc]" by auto with major have "rev() = [Oc]" by auto then have "[Oc] @ [] = rev() \ noDblBk [] \ ([] = [] \ [] \ [] \ last [] = Oc) \ () = () \ noDblBk ()" by (simp add: noDblBk_Nil noDblBk_tape_of_nat_list) then show ?thesis by blast qed lemma tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_n_eq_1_last_neq_0: assumes "(nl::nat list) \ []" and "n=1" and "n \ length nl" and "0 < last (take n nl)" shows "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" proof - have minor: "rev() = " by (rule rev_numeral_list) also with assms have "... = " by simp also have "... = <(rev[last (take n nl)]) @ (rev (butlast (take n nl)))>" by simp finally have major: "rev() = <(rev[(last (take n nl))]) @ (rev ( butlast (take n nl)))>" by auto moreover from assms have "[last (take n nl)] \ []" by auto moreover from assms have "butlast (take n nl) = []" by (simp add: butlast_take) ultimately have "rev() = <(rev[(last (take n nl))])>" by auto also have "<(rev[(last (take n nl))])> = Oc\ Suc (last (take n nl))" proof - from assms have "<[(last (take n nl))]> = Oc\ Suc (last (take n nl))" by (simp add: tape_of_list_def tape_of_nat_def) then show "<(rev[(last (take n nl))])> = Oc\ Suc (last (take n nl))" by simp qed also have "... = Oc# Oc\ (last (take n nl))" by auto finally have "rev() = Oc# Oc\ (last (take n nl))" by auto moreover from assms have "Oc\ (last (take n nl)) \ []" by auto ultimately have "[Oc] @ (Oc\ (last (take n nl))) = rev() \ noDblBk (Oc\ (last (take n nl))) \ (Oc\ (last (take n nl))) \ [] \ last (Oc\ (last (take n nl))) = Oc \ () = () \ noDblBk ()" using assms by (simp add: noDblBk_Bk_Oc_rep noDblBk_tape_of_nat_list) then show ?thesis by auto qed (* convenient versions using hd and tl for tm_skip_first_arg *) lemma tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_eq_Nil_last_eq_0': assumes "1 \ length (nl:: nat list)" and "hd nl = 0" shows "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL = [] \ CR = () \ noDblBk CR" proof - from assms have "(nl::nat list) \ [] \ (1::nat)=1 \ 1 \ length nl \ 0 = last (take 1 nl)" by (metis One_nat_def append.simps(1) append_butlast_last_id butlast_take diff_Suc_1 hd_take le_numeral_extra(4) length_0_conv less_numeral_extra(1) list.sel(1) not_one_le_zero take_eq_Nil zero_less_one) then have "\n. (nl::nat list) \ [] \ n=1 \ n \ length nl \ 0 = last (take n nl)" by blast then obtain n where w_n: "(nl::nat list) \ [] \ n=1 \ n \ length nl \ 0 = last (take n nl)" by blast then have "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL = [] \ CR = () \ noDblBk CR" using tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_eq_Nil_n_eq_1_last_eq_0 by auto then obtain CL CR where w_CL_CR: "[Oc] @ CL = rev () \ noDblBk CL \ CL = [] \ CR = \ noDblBk CR" by blast with assms w_n show ?thesis by (simp add: noDblBk_Nil noDblBk_tape_of_nat_list rev_numeral tape_of_nat_def) qed lemma tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_last_neq_0': assumes "1 \ length (nl::nat list)" and "0 < hd nl" shows "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" proof - from assms have "(nl::nat list) \ [] \ (1::nat)=1 \ 1 \ length nl \ 0 < last (take 1 nl)" by (metis append.simps(1) append_butlast_last_id butlast_take cancel_comm_monoid_add_class.diff_cancel ex_least_nat_le hd_take le_trans list.sel(1) list.size(3) neq0_conv not_less not_less_zero take_eq_Nil zero_less_one zero_neq_one) then have "\n. (nl::nat list) \ [] \ n=1 \ n \ length nl \ 0 < last (take n nl)" by blast then obtain n where w_n: "(nl::nat list) \ [] \ n=1 \ n \ length nl \ 0 < last (take n nl)" by blast then have "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" using tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_n_eq_1_last_neq_0 by auto with assms w_n show ?thesis by (simp add: drop_Suc take_Suc tape_of_list_def ) qed (*--------------------------------------------------------------------------------- *) (* generic case , where we tried to check for more than one one argument and failed *) (*--------------------------------------------------------------------------------- *) lemma tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_n_gt_1_last_eq_0: assumes "(nl::nat list) \ []" and "1 length nl" and "last (take n nl) = 0" shows "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" proof - have minor: "rev() = " by (rule rev_numeral_list) also with assms have "... = " by (metis append_butlast_last_id not_one_less_zero take_eq_Nil) also have "... = < (rev [(last (take n nl))]) @ (rev ( butlast (take n nl)))>" by simp also with assms have "... = <(rev [0]) @ (rev ( butlast (take n nl)))>" by auto finally have major: "rev() = <(rev [0]) @ (rev ( butlast (take n nl)))>" by auto moreover have "<(rev [0::nat])> = [Oc]" by (simp add: tape_of_list_def tape_of_nat_def ) moreover with assms have not_Nil: "rev (butlast (take n nl)) \ []" by (simp add: butlast_take) ultimately have "rev() = [Oc] @ [Bk] @ " using tape_of_nat_def tape_of_nat_list_cons_eq by auto then show ?thesis using major and minor and not_Nil by (metis append_Nil append_is_Nil_conv append_is_Nil_conv last_append last_appendR list.sel(3) noDblBk_tape_of_nat_list noDblBk_tape_of_nat_list_imp_noDblBk_tl numeral_list_last_is_Oc rev.simps(1) rev_append snoc_eq_iff_butlast tl_append2) qed lemma tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_n_gt_1_last_neq_0: assumes "(nl::nat list) \ []" and "1 < n" and "n \ length nl" and "0 < last (take n nl)" shows "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" proof - have minor: "rev() = " by (rule rev_numeral_list) also with assms have "... = " by (metis append_butlast_last_id not_one_less_zero take_eq_Nil) also have "... = <(rev [(last (take n nl))]) @ (rev ( butlast (take n nl)))>" by simp finally have major: "rev() = <(rev[(last (take n nl))]) @ (rev ( butlast (take n nl)))>" by auto moreover from assms have "[last (take n nl)] \ []" by auto moreover from assms have "butlast (take n nl) \ []" by (simp add: butlast_take) ultimately have "rev() = <(rev[(last (take n nl))])> @[Bk] @ <(rev ( butlast (take n nl)))>" by (metis append_numeral_list rev.simps(1) rev_rev_ident rev_singleton_conv) also have "<(rev[(last (take n nl))])> = Oc\ Suc (last (take n nl))" proof - from assms have "<[(last (take n nl))]> = Oc\ Suc (last (take n nl))" by (simp add: tape_of_list_def tape_of_nat_def) then show "<(rev[(last (take n nl))])> = Oc\ Suc (last (take n nl))" by simp qed also have "... = Oc# Oc\ (last (take n nl))" by auto finally have "rev() = Oc# Oc\ (last (take n nl)) @[Bk] @ <(rev ( butlast (take n nl)))>" by auto moreover from assms have "Oc\ (last (take n nl)) \ []" by auto ultimately have "[Oc] @ (Oc\ (last (take n nl)) @[Bk] @ <(rev ( butlast (take n nl)))>) = rev() \ noDblBk (Oc\ (last (take n nl)) @[Bk] @ <(rev ( butlast (take n nl)))>) \ (Oc\ (last (take n nl)) @[Bk] @ <(rev ( butlast (take n nl)))>) \ [] \ last (Oc\ (last (take n nl)) @[Bk] @ <(rev ( butlast (take n nl)))>) = Oc \ () = () \ noDblBk ()" using assms \ = Oc \ Suc (last (take n nl))\ \Oc \ Suc (last (take n nl)) = Oc # Oc \ last (take n nl)\ \butlast (take n nl) \ []\ \rev () = @ [Bk] @ \ - by (smt + by (smt (verit) append_Cons append_Nil append_Nil2 append_eq_Cons_conv butlast.simps(1) butlast.simps(2) butlast_append last_ConsL last_append last_appendR list.sel(3) list.simps(3) minor noDblBk_tape_of_nat_list noDblBk_tape_of_nat_list_imp_noDblBk_tl numeral_list_last_is_Oc rev_is_Nil_conv self_append_conv) then show ?thesis by auto qed lemma tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_n_gt_1: assumes "(nl::nat list) \ []" and "1 length nl" shows "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" proof (cases "last (take n nl)") case 0 with assms show ?thesis using tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_n_gt_1_last_eq_0 by auto next case (Suc nat) with assms show ?thesis using tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_n_gt_1_last_neq_0 by auto qed (*--------------------------------------------------------------------------------- *) (* For now, we only proof a result for n = 1 supporting tm_check_for_one_arg *) (*--------------------------------------------------------------------------------- *) (*----------------------------------------------------------------------------------------------------- *) (* ENHANCE: formalise generic tm_skip_n_args *) (* ENHANCE: formalise matching results for tm_erase_right_then_dblBk_left_erp_total_correctness_one_arg *) (*----------------------------------------------------------------------------------------------------- *) lemma tm_erase_right_then_dblBk_left_erp_total_correctness_one_arg: assumes "1 \ length (nl::nat list)" shows "\ \tap. tap = (Bk# rev(), ) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ () @ [Bk] @ Bk \ rex ) \" proof (cases "hd nl") case 0 then have "hd nl = 0" . with assms have "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL = [] \ CR = () \ noDblBk CR" using tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_eq_Nil_last_eq_0' by blast then obtain CL CR where w_CL_CR: "[Oc] @ CL = rev() \ noDblBk CL \ CL = [] \ CR = () \ noDblBk CR" by blast show ?thesis proof (rule Hoare_consequence) (* 1. \tap. tap = (Bk # rev (), ) \ ?P *) from assms and w_CL_CR show "(\tap. tap = (Bk # rev (), )) \ (\tap. tap = ([Bk,Oc] @ CL, CR))" using Cons_eq_appendI append_self_conv assert_imp_def by auto next (* 1. \\tap. tap = ([Bk, Oc] @ CL, CR)\ tm_erase_right_then_dblBk_left \?Q\ *) from assms and w_CL_CR show "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" using tm_erase_right_then_dblBk_left_erp_total_correctness_CL_is_Nil by blast next (* 1. \tap. \rex. tap = ([], [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex) \ \tap. \rex. tap = ([], [Bk, Bk] @ @ [Bk] @ Bk \ rex) *) show "(\tap. \rex. tap = ([], [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex)) \ (\tap. \rex. tap = ([], [Bk, Bk] @ @ [Bk] @ Bk \ rex))" using Cons_eq_append_conv assert_imp_def rev_numeral w_CL_CR by fastforce qed next case (Suc nat) then have "0 < hd nl" by auto with assms have "\CL CR. [Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" using tm_erase_right_then_dblBk_left_erp_total_correctness_helper_CL_neq_Nil_last_neq_0' by auto then obtain CL CR where w_CL_CR: "[Oc] @ CL = rev() \ noDblBk CL \ CL \ [] \ last CL = Oc \ CR = () \ noDblBk CR" by blast show ?thesis proof (rule Hoare_consequence) (* 1. \tap. tap = (Bk # rev (), ) \ ?P *) from assms and w_CL_CR show "(\tap. tap = (Bk # rev (), )) \ (\tap. tap = ([Bk,Oc] @ CL, CR))" by (simp add: w_CL_CR assert_imp_def) next (* \\tap. tap = ([Bk, Oc] @ CL, CR)\ tm_erase_right_then_dblBk_left \?Q\ *) from assms and w_CL_CR show "\ \tap. tap = ([Bk,Oc] @ CL, CR) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk, Bk] @ (rev CL) @ [Oc, Bk] @ Bk \ rex ) \" using tm_erase_right_then_dblBk_left_erp_total_correctness_CL_ew_Oc by blast next (* \tap. \rex. tap = ([], [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex) \ \tap. \rex. tap = ([], [Bk, Bk] @ @ [Bk] @ Bk \ rex) *) show "(\tap. \rex. tap = ([], [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk \ rex)) \ (\tap. \rex. tap = ([], [Bk, Bk] @ @ [Bk] @ Bk \ rex))" using Cons_eq_append_conv assert_imp_def rev_numeral w_CL_CR by (simp add: assert_imp_def rev_numeral replicate_app_Cons_same tape_of_nat_def) qed qed (* ----------------------- tm_check_for_one_arg ----------------------- * Prove total correctness for COMPOSED machine tm_check_for_one_arg * This is done via the rules for the composition operator seq_tm * -------------------------------------------------------------------- *) (* we have DO_NOTHING path corollary tm_skip_first_arg_correct_Nil': "length nl = 0 \ \\tap. tap = ([], ) \ tm_skip_first_arg \\tap. tap = ( [] , [Bk] ) \" lemma tm_skip_first_arg_len_eq_1_total_correctness': "length nl = 1 \ \\tap. tap = ([], )\ tm_skip_first_arg \ \tap. tap = ([Bk], <[hd nl]> @[Bk])\" lemma tm_erase_right_then_dblBk_left_dnp_total_correctness: "\ \tap. tap = ([], r ) \ tm_erase_right_then_dblBk_left \ \tap. tap = ([Bk,Bk], r ) \" --- ERASE path lemma tm_skip_first_arg_len_gt_1_total_correctness: assumes "1 < length (nl::nat list)" shows "\\tap. tap = ([], )\ tm_skip_first_arg \ \tap. tap = (Bk# , ) \" lemma tm_erase_right_then_dblBk_left_erp_total_correctness_one_arg: assumes "1 \ length (nl::nat list)" shows "\ \tap. tap = (Bk# rev(), ) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ () @ [Bk] @ Bk \ rex ) \" *) (* let tm_check_for_one_arg = foldl1' seqcomp_tm [ tm_skip_first_arg, tm_erase_right_then_dblBk_left ] *) definition tm_check_for_one_arg :: "instr list" where "tm_check_for_one_arg \ tm_skip_first_arg |+| tm_erase_right_then_dblBk_left" lemma tm_check_for_one_arg_total_correctness_Nil: "length nl = 0 \ \\tap. tap = ([], ) \ tm_check_for_one_arg \\tap. tap = ([Bk,Bk], [Bk] ) \" proof - assume major: "length nl = 0" have "\\tap. tap = ([], ) \ (tm_skip_first_arg |+| tm_erase_right_then_dblBk_left) \\tap. tap = ([Bk,Bk], [Bk] ) \" proof (rule Hoare_plus_halt) show "composable_tm0 tm_skip_first_arg" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_skip_first_arg_def tm_erase_right_then_dblBk_left_def) next from major show "\\tap. tap = ([], ) \ tm_skip_first_arg \\tap. tap = ( [] , [Bk] ) \" using tm_skip_first_arg_correct_Nil' by simp next from major show "\\tap. tap = ([], [Bk])\ tm_erase_right_then_dblBk_left \\tap. tap = ([Bk, Bk], [Bk])\" using tm_erase_right_then_dblBk_left_dnp_total_correctness by simp qed then show ?thesis unfolding tm_check_for_one_arg_def by auto qed lemma tm_check_for_one_arg_total_correctness_len_eq_1: "length nl = 1 \ \\tap. tap = ([], ) \ tm_check_for_one_arg \\tap. \z4. tap = (Bk \ z4, @ [Bk])\" proof - assume major: "length nl = 1" have " \\tap. tap = ([], ) \ (tm_skip_first_arg |+| tm_erase_right_then_dblBk_left) \\tap. \z4. tap = (Bk \ z4, @ [Bk])\" proof (rule Hoare_plus_halt) show "composable_tm0 tm_skip_first_arg" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_skip_first_arg_def tm_erase_right_then_dblBk_left_def) next from major have "\\tap. tap = ([], ) \ tm_skip_first_arg \ \tap. tap = ([Bk], <[hd nl]> @[Bk])\" using tm_skip_first_arg_len_eq_1_total_correctness' by simp moreover from major have "(nl::nat list) = [hd nl]" by (metis diff_self_eq_0 length_0_conv length_tl list.exhaust_sel zero_neq_one) ultimately show "\\tap. tap = ([], ) \ tm_skip_first_arg \ \tap. tap = ([Bk], @[Bk])\" using major by auto next from major have "\\tap. tap = ([], @ [Bk])\ tm_erase_right_then_dblBk_left \\tap. tap = ([Bk, Bk], @ [Bk])\" using tm_erase_right_then_dblBk_left_dnp_total_correctness by simp (* Add a blank on the initial left tape *) with major have "\stp. is_final (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, [Bk, Bk], @ [Bk]))" unfolding Hoare_halt_def - by (smt Hoare_halt_def Pair_inject holds_for.elims(2) is_final.elims(2)) + by (smt (verit) Hoare_halt_def Pair_inject holds_for.elims(2) is_final.elims(2)) then obtain stp where w_stp: "is_final (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, [Bk, Bk], @ [Bk]))" by blast then have "is_final (steps0 (1, Bk \ 0,@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, Bk \ 0,@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk \ 2, @ [Bk]))" by (simp add: is_finalI numeral_eqs_upto_12(1)) then have "\z3. z3 \ 0 + 1 \ is_final (steps0 (1, Bk \ (0+1),@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, Bk \ (0+1),@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk \ (2+z3), @ [Bk]))" by (metis is_finalI steps_left_tape_EnlargeBkCtx) then have "is_final (steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (\z4. steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk \ z4, @ [Bk]))" by (metis One_nat_def add.left_neutral replicate_0 replicate_Suc) then have "\n. is_final (steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left n) \ (\z4. steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left n = (0, Bk \ z4, @ [Bk]))" by blast then show "\\tap. tap = ([Bk], @ [Bk])\ tm_erase_right_then_dblBk_left \\tap. \z4. tap = (Bk \ z4, @ [Bk])\" using Hoare_halt_def Hoare_unhalt_def holds_for.simps by auto qed then show ?thesis unfolding tm_check_for_one_arg_def by auto qed (* Old version of tm_check_for_one_arg_total_correctness_len_eq_1: * Keep this proof as an example * The proof shows how to add blanks on the initial left tape, if you need some for composition with some predecessor tm *) (* lemma tm_check_for_one_arg_total_correctness_len_eq_1: "length nl = 1 \ \\tap. tap = ([], ) \ tm_check_for_one_arg \\tap. \z4. tap = (Bk \ z4, @ [Bk])\" proof - assume major: "length nl = 1" have " \\tap. tap = ([], ) \ (tm_skip_first_arg |+| tm_erase_right_then_dblBk_left) \\tap. \z4. tap = (Bk \ z4, @ [Bk])\" proof (rule Hoare_plus_halt) show "composable_tm0 tm_skip_first_arg" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_skip_first_arg_def tm_erase_right_then_dblBk_left_def) next from major have "\\tap. tap = ([], ) \ tm_skip_first_arg \ \tap. tap = ([Bk], <[hd nl]> @[Bk])\" using tm_skip_first_arg_len_eq_1_total_correctness' by simp moreover from major have "(nl::nat list) = [hd nl]" by (metis diff_self_eq_0 length_0_conv length_tl list.exhaust_sel zero_neq_one) ultimately show "\\tap. tap = ([], ) \ tm_skip_first_arg \ \tap. tap = ([Bk], @[Bk])\" using major by auto next from major have "\\tap. tap = ([], @ [Bk])\ tm_erase_right_then_dblBk_left \\tap. tap = ([Bk, Bk], @ [Bk])\" using tm_erase_right_then_dblBk_left_dnp_total_correctness by simp (* Add a blank on the initial left tape *) (* Note: since we proved \\tap. tap = ([], @ [Bk])\ tm_erase_right_then_dblBk_left \\tap. tap = ([Bk, Bk], @ [Bk])\ we need to add a blank on the left tape. This is no problem, since blanks on the left tape do not matter. However, adding blanks also changes the resulting left tape. Instead of \\tap. tap = ([Bk, Bk], @ [Bk])\ we end up with \\tap. \z4. tap = (Bk \ z4, @ [Bk])\ If we had proved \\tap. tap = ([Bk], @ [Bk])\ tm_erase_right_then_dblBk_left \\tap. tap = ([Bk, Bk], @ [Bk])\ in the first place, there would be no need for this weakened result. But since trailing blanks on the final left tape do not matter either, this is no real problem. *) with major have "\stp. is_final (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, [Bk, Bk], @ [Bk]))" unfolding Hoare_halt_def by (smt Hoare_halt_def Pair_inject holds_for.elims(2) is_final.elims(2)) then obtain stp where w_stp: "is_final (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, [],@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, [Bk, Bk], @ [Bk]))" by blast then have "is_final (steps0 (1, Bk \ 0,@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, Bk \ 0,@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk \ 2, @ [Bk]))" by (simp add: is_finalI numeral_eqs_upto_12(1)) then have "\z3. z3 \ 0 + 1 \ is_final (steps0 (1, Bk \ (0+1),@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (steps0 (1, Bk \ (0+1),@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk \ (2+z3), @ [Bk]))" by (metis is_finalI steps_left_tape_EnlargeBkCtx) then have "is_final (steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left stp) \ (\z4. steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk \ z4, @ [Bk]))" by (metis One_nat_def add.left_neutral replicate_0 replicate_Suc) then have "\n. is_final (steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left n) \ (\z4. steps0 (1, [Bk],@ [Bk]) tm_erase_right_then_dblBk_left n = (0, Bk \ z4, @ [Bk]))" by blast then show "\\tap. tap = ([Bk], @ [Bk])\ tm_erase_right_then_dblBk_left \\tap. \z4. tap = (Bk \ z4, @ [Bk])\" using Hoare_halt_def Hoare_unhalt_def holds_for.simps by auto qed then show ?thesis unfolding tm_check_for_one_arg_def by auto qed *) lemma tm_check_for_one_arg_total_correctness_len_gt_1: "length nl > 1 \ \\tap. tap = ([], )\ tm_check_for_one_arg \ \tap. \ l. tap = ( [], [Bk,Bk] @ <[hd nl]> @ Bk \ l) \" proof - assume major: "length nl > 1" have " \\tap. tap = ([], )\ (tm_skip_first_arg |+| tm_erase_right_then_dblBk_left) \ \tap. \ l. tap = ( [], [Bk,Bk] @ <[hd nl]> @ Bk \ l) \" proof (rule Hoare_plus_halt) show "composable_tm0 tm_skip_first_arg" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_skip_first_arg_def tm_erase_right_then_dblBk_left_def) next from major show "\\tap. tap = ([], )\ tm_skip_first_arg \ \tap. tap = (Bk# , ) \" using tm_skip_first_arg_len_gt_1_total_correctness by simp next from major have "\ \tap. tap = (Bk# rev(), ) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ () @ [Bk] @ Bk \ rex ) \" using tm_erase_right_then_dblBk_left_erp_total_correctness_one_arg by simp then have "\ \tap. tap = (Bk# , ) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ <[hd nl]> @ [Bk] @ Bk \ rex ) \" by (simp add: rev_numeral rev_numeral_list tape_of_list_def ) then have "\ \tap. tap = (Bk# , ) \ tm_erase_right_then_dblBk_left \ \tap. \rex. tap = ([], [Bk,Bk] @ <[hd nl]> @ Bk \ (Suc rex) ) \" by force then show "\\tap. tap = (Bk # , )\ tm_erase_right_then_dblBk_left \\tap. \l. tap = ([], [Bk, Bk] @ <[hd nl]> @ Bk \ l)\" - by (smt Hoare_haltI Hoare_halt_def holds_for.elims(2) holds_for.simps) + by (smt (verit) Hoare_haltI Hoare_halt_def holds_for.elims(2) holds_for.simps) qed then show ?thesis unfolding tm_check_for_one_arg_def by auto qed (* ---------- tm_strong_copy ---------- *) definition tm_strong_copy :: "instr list" where "tm_strong_copy \ tm_check_for_one_arg |+| tm_weak_copy" (* Prove total correctness for COMPOSED machine tm_strong_copy. * This is done via the rules for the composition operator seq_tm. *) lemma tm_strong_copy_total_correctness_Nil: "length nl = 0 \ \\tap. tap = ([], ) \ tm_strong_copy \\tap. tap = ([Bk,Bk,Bk,Bk],[]) \" proof - assume major: "length nl = 0" have " \\tap. tap = ([], ) \ tm_check_for_one_arg |+| tm_weak_copy \\tap. tap = ([Bk,Bk,Bk,Bk],[]) \" proof (rule Hoare_plus_halt) show "composable_tm0 tm_check_for_one_arg" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_weak_copy_def tm_check_for_one_arg_def tm_skip_first_arg_def tm_erase_right_then_dblBk_left_def) next from major show "\\tap. tap = ([], ) \ tm_check_for_one_arg \\tap. tap = ([Bk,Bk], [Bk] ) \" using tm_check_for_one_arg_total_correctness_Nil by simp next from major show "\\tap. tap = ([Bk, Bk], [Bk])\ tm_weak_copy \\tap. tap = ([Bk, Bk, Bk, Bk], [])\" using tm_weak_copy_correct11' by simp qed then show ?thesis unfolding tm_strong_copy_def by auto qed lemma tm_strong_copy_total_correctness_len_gt_1: "1 < length nl \ \\tap. tap = ([], ) \ tm_strong_copy \ \tap. \l. tap = ([Bk,Bk], <[hd nl]> @ Bk \ l) \" proof - assume major: "1 < length nl" have " \\tap. tap = ([], ) \ tm_check_for_one_arg |+| tm_weak_copy \ \tap. \ l. tap = ([Bk,Bk], <[hd nl]> @ Bk \ l) \" proof (rule Hoare_plus_halt) show "composable_tm0 tm_check_for_one_arg" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_weak_copy_def tm_check_for_one_arg_def tm_skip_first_arg_def tm_erase_right_then_dblBk_left_def) next from major show "\\tap. tap = ([], ) \ tm_check_for_one_arg \ \tap. \ l. tap = ( [], [Bk,Bk] @ <[hd nl]> @ Bk \ l) \" using tm_check_for_one_arg_total_correctness_len_gt_1 by simp next show "\\tap. \l. tap = ([], [Bk, Bk] @ <[hd nl]> @ Bk \ l)\ tm_weak_copy \\tap. \l. tap = ([Bk, Bk], <[hd nl]> @ Bk \ l)\" proof - have "\r.\\tap. tap = ([], [Bk,Bk]@r) \ tm_weak_copy \\tap. tap = ([Bk,Bk], r) \" using tm_weak_copy_correct13' by simp then have "\r. \stp. is_final (steps0 (1, [],[Bk,Bk]@r) tm_weak_copy stp) \ (steps0 (1, [],[Bk,Bk]@r) tm_weak_copy stp = (0, [Bk, Bk], r))" unfolding Hoare_halt_def - by (smt Hoare_halt_def Pair_inject holds_for.elims(2) is_final.elims(2)) + by (smt (verit) Hoare_halt_def Pair_inject holds_for.elims(2) is_final.elims(2)) then have "\l. \stp. is_final (steps0 (1, [],[Bk,Bk]@<[hd nl]> @ Bk \ l) tm_weak_copy stp) \ (steps0 (1, [],[Bk,Bk]@<[hd nl]> @ Bk \ l) tm_weak_copy stp = (0, [Bk, Bk], <[hd nl]> @ Bk \ l))" by blast then show ?thesis using Hoare_halt_def holds_for.simps by fastforce qed qed then show ?thesis unfolding tm_strong_copy_def by auto qed lemma tm_strong_copy_total_correctness_len_eq_1: "1 = length nl \ \\tap. tap = ([], ) \ tm_strong_copy \ \tap. \k l. tap = (Bk \ k, <[hd nl, hd nl]> @ Bk \ l) \" proof - assume major: "1 = length nl" have " \\tap. tap = ([], ) \ tm_check_for_one_arg |+| tm_weak_copy \ \tap. \k l. tap = (Bk \ k, <[hd nl, hd nl]> @ Bk \ l) \" proof (rule Hoare_plus_halt) show "composable_tm0 tm_check_for_one_arg" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_weak_copy_def tm_check_for_one_arg_def tm_skip_first_arg_def tm_erase_right_then_dblBk_left_def) next from major show "\\tap. tap = ([], ) \ tm_check_for_one_arg \\tap. \z4. tap = (Bk \ z4, @ [Bk])\" using tm_check_for_one_arg_total_correctness_len_eq_1 by simp next have "\\tap. \z4. tap = (Bk \ z4, <[hd nl]> @[Bk])\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, <[hd nl, hd nl]> @ Bk \ l) \" using tm_weak_copy_correct6 by simp moreover from major have " = <[hd nl]>" by (metis diff_self_eq_0 length_0_conv length_tl list.exhaust_sel zero_neq_one) ultimately show "\\tap. \z4. tap = (Bk \ z4, @ [Bk])\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, <[hd nl, hd nl]> @ Bk \ l)\" by auto qed then show ?thesis unfolding tm_strong_copy_def by auto qed end diff --git a/thys/Universal_Turing_Machine/TM_Playground.thy b/thys/Universal_Turing_Machine/TM_Playground.thy --- a/thys/Universal_Turing_Machine/TM_Playground.thy +++ b/thys/Universal_Turing_Machine/TM_Playground.thy @@ -1,97 +1,97 @@ (* Title: thys/TM_Playground.thy Author: Franz Regensburger (FABR) 02/2022 *) theory TM_Playground imports StrongCopyTM begin (* Playground *) (* ad numerals *) value "[1::nat,2,3,4] ! 0" value "[1::nat,2,3,4] ! 1" value "<2::nat>" value "tape_of (2::nat)" value "<(0::nat,1::nat)>" value "tape_of_nat_prod (0::nat,1::nat)" value "<([0,1::nat],[0,1::nat])>" value "tape_of_nat_prod ([0,1::nat],[0,1::nat])" value "tape_of_nat_list [2,0,3::nat]" value "<[2,0,3::nat]>" (* ad fetch function *) lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 0 Bk = (Nop , 0)" by auto lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 0 Oc = (Nop , 0)" by auto lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 1 Bk = (R , 3)" by auto lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 1 Oc = (WB, 2)" by auto lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 3 Bk = (WO, 0)" by (simp add: numeral_eqs_upto_12) lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 3 Oc = (WB, 2)" by (simp add: numeral_eqs_upto_12) (* fetch last instruction of program, which is explicitly given by the instr list *) lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 4 Bk = (L, 8)" by (simp add: numeral_eqs_upto_12) (* fetch instruction outside of explicitly given instr list *) lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 4 Oc = (Nop, 0)" by (simp add: numeral_eqs_upto_12) lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 10 Bk = (Nop, 0)" by (simp add: numeral_eqs_upto_12) lemma "fetch [(R, 3),(WB,2), (R,1),(R,1), (WO,0),(WB,2), (L,8)] 10 Oc = (Nop, 0)" by (simp add: numeral_eqs_upto_12) (* ad Hoare triples *) (* This result can be validated by a simple test run *) lemma "\\(l1, r1). (l1,r1) = ([], [Oc,Oc] )\ tm_weak_copy \\(l0,r0). \k l. (l0,r0) = (Bk \ k, [Oc,Oc,Bk,Oc,Oc] @ Bk \ l)\" proof - have "\\tap. tap = ([], <[1::nat]>) \ tm_weak_copy \ \tap. \k l. tap = (Bk \ k, <[hd [1::nat], hd [1::nat]]> @ Bk \ l) \" by (simp add: tm_weak_copy_correct5) then have "\\tap. tap = ([], <[1::nat]>) \ tm_weak_copy \ \tap. \k l. tap = (Bk \ k, <[1::nat, 1::nat]> @ Bk \ l) \" by force moreover have "<[1::nat]> = [Oc,Oc]" by (simp add: tape_of_list_def tape_of_nat_def) moreover have "<[1::nat, 1::nat]> = [Oc,Oc,Bk,Oc,Oc]" by (simp add: tape_of_list_def tape_of_nat_def) ultimately have "\\tap. tap = ([], <[1::nat]>) \ tm_weak_copy \ \tap. \k l. tap = (Bk \ k, <[hd [1::nat], hd [1::nat]]> @ Bk \ l) \" by simp then show ?thesis - by (smt Hoare_haltI Hoare_halt_E0 \<[1, 1]> = [Oc, Oc, Bk, Oc, Oc]\ \<[1]> = [Oc, Oc]\ + by (smt (verit) Hoare_haltI Hoare_halt_E0 \<[1, 1]> = [Oc, Oc, Bk, Oc, Oc]\ \<[1]> = [Oc, Oc]\ case_prodD case_prodI holds_for.simps is_final_eq list.discI list.inject list.sel(1) not_Cons_self2 tape_of_list_def tape_of_nat_def tape_of_nat_list.elims tape_of_nat_list.simps(3)) qed (* However, this general result can never be validated by testing. It must be verified by a proof *) lemma "\n. \\tap. tap = ([], Oc\(n+1))\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, Oc\(n+1) @ [Bk] @ Oc\(n+1) @ Bk \ l) \" proof fix n have "\\tap. tap = ([], <[n::nat]>)\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, <[n, n]> @ Bk \ l) \" by (rule tm_weak_copy_correct5) moreover have "<[n::nat]> = Oc\(Suc n)" by (simp add: tape_of_list_def tape_of_nat_def) moreover have "<[n, n]> = Oc\(Suc n) @ [Bk] @ Oc\(Suc n)" by (simp add: tape_of_list_def tape_of_nat_def) ultimately have X: "\\tap. tap = ([], Oc\(Suc n))\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, Oc\(Suc n) @ [Bk] @ Oc\(Suc n) @ Bk \ l) \" by simp show "\\tap. tap = ([], Oc\(n+1))\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, Oc\(n+1) @ [Bk] @ Oc\(n+1) @ Bk \ l) \" using X by auto qed end diff --git a/thys/Universal_Turing_Machine/TuringComputable.thy b/thys/Universal_Turing_Machine/TuringComputable.thy --- a/thys/Universal_Turing_Machine/TuringComputable.thy +++ b/thys/Universal_Turing_Machine/TuringComputable.thy @@ -1,699 +1,699 @@ (* Title: thys/TuringComputable.thy Author: Franz Regensburger (FABR) 03/2022 *) section \Turing Computable Functions\ theory TuringComputable imports "HaltingProblems_K_H" begin (* -------------------------------------------------------------------------- *) (* declare adjust.simps[simp del] declare seq_tm.simps [simp del] declare shift.simps[simp del] declare composable_tm.simps[simp del] declare step.simps[simp del] declare steps.simps[simp del] *) subsection \Definition of Partial Turing Computability\ text \We present two variants for a definition of Partial Turing Computability, which we prove to be equivalent, later on.\ subsubsection \Definition Variant 1\ definition turing_computable_partial :: "(nat list \ nat option) \ bool" where "turing_computable_partial f \ (\tm. \ns n. (f ns = Some n \ (\stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \))" (* Major consequences of the definition *) (* This is for documentation and explanation: turing_computable_partial_unfolded_into_TMC_yields_TMC_has_conditions *) lemma turing_computable_partial_unfolded_into_TMC_yields_TMC_has_conditions: "turing_computable_partial f \ (\tm. \ns n. (f ns = Some n \ TMC_yields_num_res tm ns n) \ (f ns = None \ \ TMC_has_num_res tm ns ) )" unfolding TMC_yields_num_res_def TMC_has_num_res_def by (simp add: turing_computable_partial_def) (* This is for rewriting *) lemma turing_computable_partial_unfolded_into_Hoare_halt_conditions: "turing_computable_partial f \ (\tm. \ns n. (f ns = Some n \ \ \tap. tap = ([], ) \ tm \ \tap. \k l. tap = (Bk \ k, @ Bk\ l) \ ) \ (f ns = None \ \\ \tap. tap = ([], ) \ tm \ \tap. \k n l. tap = (Bk \ k, @ Bk \ l) \))" unfolding turing_computable_partial_def proof assume "\tm. \ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. \k n l. tap = (Bk \ k, @ Bk \ l) \)" then obtain tm where w_tm1:"\ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. \k n l. tap = (Bk \ k, @ Bk \ l) \)" by blast have "\ns n. (f ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. \k n l. tap = (Bk \ k, @ Bk \ l) \)" apply (safe) proof - (* ---- Some case *) show "\ns n. f ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" proof - fix ns n assume "f ns = Some n" with w_tm1 have F1: "\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" by auto show "\\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" proof (rule Hoare_haltI) fix l r assume "(l, r) = ([]::cell list, )" then show "\stp. is_final (steps0 (1, l, r) tm stp) \ (\tap. \k l. tap = (Bk \ k, @ Bk \ l)) holds_for steps0 (1, l, r) tm stp" by (metis (mono_tags, lifting) F1 holds_for.simps is_finalI) qed qed next (* ---- None case *) show "\ns n. \f ns = None; \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\ \ \ False" proof - fix ns n assume "f ns = None" and "\\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" with w_tm1 show False by simp qed qed then show "\tm. \ns n. (f ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (f ns = None \ \ \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\)" by auto next assume "\tm. \ns n. (f ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (f ns = None \ \ \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\)" then obtain tm where w_tm2: "\ns n. (f ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (f ns = None \ \ \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\)" by blast have "\ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\)" apply (safe) proof - (* ---- Some case *) show "\ns n . f ns = Some n \ \stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" proof - fix ns n z assume "f ns = Some n" with w_tm2 have "\\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" by blast then show "\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" - by (smt Hoare_halt_def holds_for.elims(2) is_final.elims(2) snd_conv) + by (smt (verit) Hoare_halt_def holds_for.elims(2) is_final.elims(2) snd_conv) qed next (* ---- None case *) show "\ns n. \f ns = None; \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\\ \ False" proof - fix ns n assume "f ns = None" and "\\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" with w_tm2 show False by simp qed qed then show "\tm. \ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\)" by blast qed (* --------------------------------------------------------------------------------------------------------- *) subsubsection \Characteristic Functions of Sets\ definition chi_fun :: "(nat list) set \ (nat list \ nat option)" where "chi_fun nls = (\nl. if nl \ nls then Some 1 else Some 0)" lemma chi_fun_0_iff: "nl \ nls \ chi_fun nls nl = Some 0" unfolding chi_fun_def by auto lemma chi_fun_1_iff: "nl \ nls \ chi_fun nls nl = Some 1" unfolding chi_fun_def by auto lemma chi_fun_0_I: "nl \ nls \ chi_fun nls nl = Some 0" unfolding chi_fun_def by auto lemma chi_fun_0_E: "(chi_fun nls nl = Some 0 \ P) \ nl \ nls \ P" unfolding chi_fun_def by auto lemma chi_fun_1_I: "nl \ nls \ chi_fun nls nl = Some 1" unfolding chi_fun_def by auto lemma chi_fun_1_E: "(chi_fun nls nl = Some 1 \ P) \ nl \ nls \ P" unfolding chi_fun_def by auto (* --------------------------------------------------------------------------------------------------------- *) subsubsection \Relation between Partial Turing Computability and Turing Decidability\ text \If a set A is Turing Decidable its characteristic function is Turing Computable partial and vice versa. Please note, that although the characteristic function has an option type it will always yield Some value. \ theorem turing_decidable_imp_turing_computable_partial: "turing_decidable A \ turing_computable_partial (chi_fun A)" proof - assume "turing_decidable A" then have "(\D. (\nl. (nl \ A \ \(\tap. tap = ([], ))\ D \(\tap. \k l. tap = (Bk \ k, <1::nat> @ Bk\ l))\) \ (nl \ A \ \(\tap. tap = ([], ))\ D \(\tap. \k l. tap = (Bk \ k, <0::nat> @ Bk\ l))\) ))" unfolding turing_decidable_def by auto then obtain D where w_D: "(\nl. (nl \ A \ \(\tap. tap = ([], ))\ D \(\tap. \k l. tap = (Bk \ k, <1::nat> @ Bk\ l))\) \ (nl \ A \ \(\tap. tap = ([], ))\ D \(\tap. \k l. tap = (Bk \ k, <0::nat> @ Bk\ l))\) )" by blast then have F1: "(\nl. (chi_fun A nl = Some 1 \ \(\tap. tap = ([], ))\ D \(\tap. \k l. tap = (Bk \ k, <1::nat> @ Bk\ l))\) \ (chi_fun A nl = Some 0 \ \(\tap. tap = ([], ))\ D \(\tap. \k l. tap = (Bk \ k, <0::nat> @ Bk\ l))\) )" using chi_fun_def by force have F2: "\nl. (chi_fun A nl = Some 1 \ chi_fun A nl = Some 0)" using chi_fun_def by simp show ?thesis proof (rule turing_computable_partial_unfolded_into_Hoare_halt_conditions[THEN iffD2]) have "\ns n. (chi_fun A ns = Some n \ \\tap. tap = ([], )\ D \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (chi_fun A ns = None \ \ \\tap. tap = ([], )\ D \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\)" apply (safe) proof - show "\ns n. chi_fun A ns = Some n \ \\tap. tap = ([], )\ D \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" proof - fix ns :: "nat list" and n :: nat assume "chi_fun A ns = Some n" then have "Some n = Some 1 \ Some n = Some 0" by (metis F2) then show "\\tap. tap = ([], )\ D \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" using \chi_fun A ns = Some n\ F1 by blast qed next show " \ns n. \chi_fun A ns = None; \\tap. tap = ([], )\ D \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\ \ \ False" by (metis F2 option.distinct(1)) qed then show "\tm. \ns n. (chi_fun A ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (chi_fun A ns = None \ \ \\tap. tap = ([], )\ tm \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\)" using F2 option.simps(3) by blast qed qed theorem turing_computable_partial_imp_turing_decidable: "turing_computable_partial (chi_fun A) \ turing_decidable A" proof - assume "turing_computable_partial (chi_fun A)" then have "\tm. \ns n. (chi_fun A ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (chi_fun A ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \ )" using turing_computable_partial_unfolded_into_Hoare_halt_conditions[THEN iffD1] by auto then obtain tm where w_tm: "\ns n. (chi_fun A ns = Some n \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\) \ (chi_fun A ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" by auto then have "\ns. (ns \ A \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, <1::nat> @ Bk \ l)\) \ (ns \ A \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, <0::nat> @ Bk \ l)\)" by (blast intro: chi_fun_0_I chi_fun_1_I) then have "(\D.(\ns. (ns \ A \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, <1::nat> @ Bk \ l)\) \ (ns \ A \ \\tap. tap = ([], )\ tm \\tap. \k l. tap = (Bk \ k, <0::nat> @ Bk \ l)\) ))" by auto then show ?thesis using turing_decidable_def by auto qed corollary turing_computable_partial_iff_turing_decidable: "turing_decidable A \ turing_computable_partial (chi_fun A)" by (auto simp add: turing_computable_partial_imp_turing_decidable turing_decidable_imp_turing_computable_partial) subsubsection \Examples for uncomputable functions\ text \Now, we prove that the characteristic functions of the undecidable sets K1 and H1 are both uncomputable.\ context hpk begin theorem "\(turing_computable_partial (chi_fun K1))" using not_Turing_decidable_K1 turing_computable_partial_imp_turing_decidable by blast theorem "\(turing_computable_partial (chi_fun H1))" using not_Turing_decidable_H1 turing_computable_partial_imp_turing_decidable by blast end subsubsection \The Function associated with a Turing Machine\ text \With every Turing machine, we can associate a function.\ (* Compare to definition of \\<^sub>P(r1, ... , rm) in the book [DSW94, pp 29] Computability, Complexity, and Languages Davis, Sigal and Weyuker Academic Press, 2nd Edition, 1994 ISBN 0-12-206382-1 Relation between our notion fun_of_tm and \\<^sub>P of [DSW94, pp 29] \\<^sub>P(r1, ... , rm) = y \ fun_of_tm tm <[r1, ... , rm]> = Some y and \\<^sub>P(r1, ... , rm)\ \ fun_of_tm tm <[r1, ... , rm]> = None Note: for the GOTO pgms of [DSW94] the question whether the program P halts with a standard tape or with a non-standard tape does not matter since there are no tapes to consider. There are only registers with values. If the program P halts on input (r1, ... , rm) it yields a result, always. On the other hand, our Turing machines may reach the final state 0 with a non-standard tape and thus provide no valid output. In such a case, fun_of_tm tm <[r1, ... , rm]> = None per definition. *) definition fun_of_tm :: "tprog0 \ (nat list \ nat option)" where "fun_of_tm tm ns \ (if \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \ then let result = (THE n. \stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l)) in Some result else None)" text \Some immediate consequences of the definition.\ (* This is for documentation and explanation: fun_of_tm_unfolded_into_TMC_yields_TMC_has_conditions *) lemma fun_of_tm_unfolded_into_TMC_yields_TMC_has_conditions: "fun_of_tm tm \ (\ns. (if TMC_has_num_res tm ns then let result = (THE n. TMC_yields_num_res tm ns n) in Some result else None) )" unfolding TMC_yields_num_res_def TMC_has_num_res_def using fun_of_tm_def by presburger lemma fun_of_tm_is_None: assumes "\(\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" shows "fun_of_tm tm ns = None" proof - from assms show "fun_of_tm tm ns = None" by (auto simp add: fun_of_tm_def) qed lemma fun_of_tm_is_None_rev: assumes "fun_of_tm tm ns = None" shows "\(\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" using assms fun_of_tm_def by auto corollary fun_of_tm_is_None_iff: "fun_of_tm tm ns = None \ \(\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" by (auto simp add: fun_of_tm_is_None fun_of_tm_is_None_rev) (* This is for documentation and explanation: fun_of_tm_is_None_iff' *) corollary fun_of_tm_is_None_iff': "fun_of_tm tm ns = None \ \ TMC_has_num_res tm ns" unfolding TMC_has_num_res_def by (auto simp add: fun_of_tm_is_None fun_of_tm_is_None_rev ) lemma fun_of_tm_ex_Some_n': assumes "\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" shows "\n. fun_of_tm tm ns = Some n" using assms fun_of_tm_def by auto lemma fun_of_tm_ex_Some_n'_rev: assumes "\n. fun_of_tm tm ns = Some n" shows "\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" using assms fun_of_tm_is_None by fastforce corollary fun_of_tm_ex_Some_n'_iff: "(\n. fun_of_tm tm ns = Some n) \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" by (auto simp add: fun_of_tm_ex_Some_n' fun_of_tm_ex_Some_n'_rev) subsubsection \Stronger results about uniqueness of results\ corollary Hoare_halt_on_numeral_list_yields_unique_list_result_iff: "\\tap. tap = ([], )\ p \\tap. \kr ml lr. tap = (Bk \ kr, @ Bk \ lr)\ \ (\!ml. \stp k l. steps0 (1,[], ) p stp = (0, Bk \ k, @ Bk \ l))" proof assume A: "\\tap. tap = ([], )\ p \\tap. \kr ml lr. tap = (Bk \ kr, @ Bk \ lr)\" show "\!ml. \stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" proof (rule ex_ex1I) show "\ml stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" using A using Hoare_halt_iff by auto next show "\ml y. \\stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l); \stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)\ \ ml = y" by (metis nat_le_linear prod.inject prod.inject stable_config_after_final_ge' unique_Bk_postfix_numeral_list) qed next assume "(\!ml. \stp k l. steps0 (1,[], ) p stp = (0, Bk \ k, @ Bk \ l))" then show "\\tap. tap = ([], )\ p \\tap. \kr ml lr. tap = (Bk \ kr, @ Bk \ lr)\" using Hoare_halt_on_numeral_imp_list_result_rev by blast qed corollary Hoare_halt_on_numeral_yields_unique_result_iff: "\(\tap. tap = ([], ))\ p \(\tap. (\k n l. tap = (Bk \ k, @ Bk \ l)))\ \ (\!n. \stp k l. steps0 (1,[], ) p stp = (0, Bk \ k, @ Bk \ l))" proof assume A: "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" show "\!n. \stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" proof (rule ex_ex1I) show "\n stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" using A using Hoare_halt_on_numeral_imp_result by blast next show "\n y. \\stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l); \stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)\ \ n = y" - by (smt before_final is_final_eq le_less least_steps less_Suc_eq not_less_iff_gr_or_eq snd_conv unique_decomp_std_tap) + by (smt (verit) before_final is_final_eq le_less least_steps less_Suc_eq not_less_iff_gr_or_eq snd_conv unique_decomp_std_tap) qed next assume "\!n. \stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" then show "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" using Hoare_halt_on_numeral_imp_result_rev by blast qed (* --- *) lemma fun_of_tm_is_Some_unique_value: assumes "steps0 (1, ([], )) tm stp = (0, Bk \ k1, @ Bk \ l1)" shows "fun_of_tm tm ns = Some n" proof - from assms have F0: "TMC_has_num_res tm ns" using Hoare_halt_on_numeral_imp_result_rev TMC_has_num_res_def by blast then have "fun_of_tm tm ns = ( let result = (THE n. \stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l)) in Some result)" by (simp add: F0 TMC_has_num_res_def fun_of_tm_def fun_of_tm_ex_Some_n' ) then have F1: "fun_of_tm tm ns = Some (THE n. \stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l))" by auto have "(THE n. \stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l)) = n" proof (rule the1I2) from F0 have F2: "(\!n.\stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l))" using Hoare_halt_on_numeral_imp_result_rev Hoare_halt_on_numeral_yields_unique_result_iff assms by blast then show "\!n. \stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" by auto next show "\x. \stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l) \ x = n" using Hoare_halt_on_numeral_imp_result_rev Hoare_halt_on_numeral_yields_unique_result_iff assms by blast qed then show ?thesis using F1 by blast qed (* --- *) lemma fun_of_tm_ex_Some_n: assumes "\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" shows "\stp k n l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l) \ fun_of_tm tm ns = Some (n::nat)" using Hoare_halt_on_numeral_imp_result assms fun_of_tm_is_Some_unique_value by blast lemma fun_of_tm_ex_Some_n_rev: assumes "\stp k n l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l) \ fun_of_tm tm ns = Some n" shows "\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" proof - from assms have "\stp k n l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" by blast then show ?thesis using Hoare_haltE Hoare_haltI assms fun_of_tm_ex_Some_n'_rev steps.simps(1) by auto qed corollary fun_of_tm_ex_Some_n_iff: "(\stp k n l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l) \ fun_of_tm tm ns = Some n) \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" using fun_of_tm_ex_Some_n fun_of_tm_ex_Some_n_rev by blast (* --- *) lemma fun_of_tm_eq_Some_n_imp_same_numeral_result: assumes "fun_of_tm tm ns = Some n" shows "\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" by (metis (no_types, lifting) assms assms fun_of_tm_def fun_of_tm_ex_Some_n_iff fun_of_tm_is_None option.inject option.simps(3)) lemma numeral_result_n_imp_fun_of_tm_eq_n: assumes "\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" shows "fun_of_tm tm ns = Some n" using assms fun_of_tm_is_Some_unique_value by blast corollary numeral_result_n_iff_fun_of_tm_eq_n: "fun_of_tm tm ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))" using fun_of_tm_eq_Some_n_imp_same_numeral_result numeral_result_n_imp_fun_of_tm_eq_n by blast (* This is for documentation and explanation: numeral_result_n_iff_fun_of_tm_eq_n' *) corollary numeral_result_n_iff_fun_of_tm_eq_n': "fun_of_tm tm ns = Some n \ TMC_yields_num_res tm ns n" using fun_of_tm_eq_Some_n_imp_same_numeral_result numeral_result_n_imp_fun_of_tm_eq_n unfolding TMC_yields_num_res_def by blast subsubsection \Definition of Turing computability Variant 2\ definition turing_computable_partial' :: "(nat list \ nat option) \ bool" where "turing_computable_partial' f \ \tm. fun_of_tm tm = f" lemma turing_computable_partial'I: "(\ns. fun_of_tm tm ns = f ns) \ turing_computable_partial' f" unfolding turing_computable_partial'_def proof - assume "(\ns. fun_of_tm tm ns = f ns)" then have "fun_of_tm tm = f" by (rule ext) then show "\tm. fun_of_tm tm = f" by auto qed subsubsection \Definitional Variants 1 and 2 of Partial Turing Computability are equivalent\ text \Now, we prove the equivalence of the two definitions of Partial Turing Computability.\ lemma turing_computable_partial'_imp_turing_computable_partial: "turing_computable_partial' f \ turing_computable_partial f" unfolding turing_computable_partial'_def turing_computable_partial_def proof assume "\tm. fun_of_tm tm = f" then obtain tm where w_tm: "fun_of_tm tm = f" by blast show "\tm. \ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" proof show "\ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" proof fix ns show "\n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" proof fix n::nat show "(f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" proof show "f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))" proof assume "f ns = Some n" with w_tm have A: "fun_of_tm tm ns = Some n" by auto then have "\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" using fun_of_tm_ex_Some_n'_rev by auto then have "\stp k m l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l) \ fun_of_tm tm ns = Some m" by (rule fun_of_tm_ex_Some_n) with A show "\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" by auto qed next show "f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" proof assume "f ns = None" with w_tm have A: "fun_of_tm tm ns = None" by auto then show "\(\ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" by (rule fun_of_tm_is_None_rev) qed qed qed qed qed qed lemma turing_computable_partial_imp_turing_computable_partial': "turing_computable_partial f \ turing_computable_partial' f" unfolding turing_computable_partial_def proof assume major: "\tm. \ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" show "turing_computable_partial' f" proof - from major obtain tm where w_tm: "\ns n. (f ns = Some n \ (\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))) \ (f ns = None \ \ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \)" by blast show "turing_computable_partial' f" proof (rule turing_computable_partial'I) show "\ns. fun_of_tm tm ns = f ns" proof - fix ns show "fun_of_tm tm ns = f ns" proof (cases "f ns") case None then have "f ns = None" . with w_tm have "\ \ \tap. tap = ([], ) \ tm \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" by auto then have "fun_of_tm tm ns = None" by (rule fun_of_tm_is_None) with \f ns = None\ show ?thesis by auto next case (Some m) then have "f ns = Some m" . with w_tm have B: "(\stp k l. steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l))" by auto then obtain stp k l where w_stp_k_l: "steps0 (1, [], ) tm stp = (0, Bk \ k, @ Bk \ l)" by blast then have "\stp k n l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l)" by blast from this and w_stp_k_l have "fun_of_tm tm ns = Some m" by (simp add: w_stp_k_l fun_of_tm_is_Some_unique_value) with \f ns = Some m\ show ?thesis by auto qed qed qed qed qed corollary turing_computable_partial'_turing_computable_partial_iff: "turing_computable_partial' f \ turing_computable_partial f" by (auto simp add: turing_computable_partial'_imp_turing_computable_partial turing_computable_partial_imp_turing_computable_partial') text \As a now trivial consequence we obtain:\ corollary "turing_computable_partial f \ \tm. fun_of_tm tm = f" using turing_computable_partial'_turing_computable_partial_iff turing_computable_partial'_def by auto subsection \Definition of Total Turing Computability\ definition turing_computable_total :: "(nat list \ nat option) \ bool" where "turing_computable_total f \ (\tm. \ns. \n. f ns = Some n \ (\stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l)) )" (* This is for documentation and explanation: turing_computable_total_unfolded_into_TMC_yields_condition *) lemma turing_computable_total_unfolded_into_TMC_yields_condition: "turing_computable_total f \ (\tm. \ns. \n. f ns = Some n \ TMC_yields_num_res tm ns n )" unfolding TMC_yields_num_res_def by (simp add: turing_computable_total_def) (* Major consequences of the definition *) lemma turing_computable_total_imp_turing_computable_partial: "turing_computable_total f \ turing_computable_partial f" unfolding turing_computable_total_def turing_computable_partial_def by (metis option.inject option.simps(3)) corollary turing_decidable_imp_turing_computable_total_chi_fun: "turing_decidable A \ turing_computable_total (chi_fun A)" unfolding turing_computable_total_unfolded_into_TMC_yields_condition proof - assume "turing_decidable A" then have "\D. (\nl. (nl \ A \ TMC_yields_num_res D nl (1::nat) ) \ (nl \ A \ TMC_yields_num_res D nl (0::nat) ) )" unfolding turing_decidable_unfolded_into_TMC_yields_conditions by auto then obtain D where w_D: "\nl. (nl \ A \ TMC_yields_num_res D nl (1::nat) ) \ (nl \ A \ TMC_yields_num_res D nl (0::nat) )" by blast then have "\ns. \n. chi_fun A ns = Some n \ TMC_yields_num_res D ns n" by (simp add: w_D chi_fun_0_E chi_fun_def) then show "\tm. \ns. \n. chi_fun A ns = Some n \ TMC_yields_num_res tm ns n" by auto qed (* --- alternative definition turing_computable_total' --- *) definition turing_computable_total' :: "(nat list \ nat option) \ bool" where "turing_computable_total' f \ (\tm. \ns. \n. f ns = Some n \ fun_of_tm tm = f)" theorem turing_computable_total'_eq_turing_computable_total: "turing_computable_total' f = turing_computable_total f" proof show "turing_computable_total' f \ turing_computable_total f" unfolding turing_computable_total_def unfolding turing_computable_total'_def using fun_of_tm_eq_Some_n_imp_same_numeral_result by blast next show "turing_computable_total f \ turing_computable_total' f" unfolding turing_computable_total_def unfolding turing_computable_total'_def by (metis numeral_result_n_imp_fun_of_tm_eq_n tape_of_list_def turing_computable_partial'I turing_computable_partial'_def) qed (* --- alternative definition turing_computable_total'' --- *) definition turing_computable_total'' :: "(nat list \ nat option) \ bool" where "turing_computable_total'' f \ (\tm. fun_of_tm tm = f \ (\ns. \n. f ns = Some n))" theorem turing_computable_total''_eq_turing_computable_total: "turing_computable_total'' f = turing_computable_total f" proof show "turing_computable_total'' f \ turing_computable_total f" unfolding turing_computable_total_def unfolding turing_computable_total''_def by (meson fun_of_tm_eq_Some_n_imp_same_numeral_result) next show "turing_computable_total f \ turing_computable_total'' f" unfolding turing_computable_total_def unfolding turing_computable_total''_def by (metis numeral_result_n_imp_fun_of_tm_eq_n tape_of_list_def turing_computable_partial'I turing_computable_partial'_def) qed (* --- Definition for turing_computable_total_on --- *) definition turing_computable_total_on :: "(nat list \ nat option) \ (nat list) set \ bool" where "turing_computable_total_on f A \ (\tm. \ns. ns \ A \ (\n. f ns = Some n \ (\stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l))) )" (* This is for documentation and explanation: turing_computable_total_on_unfolded_into_TMC_yields_condition *) lemma turing_computable_total_on_unfolded_into_TMC_yields_condition: "turing_computable_total_on f A \ (\tm. \ns. ns \ A \ (\n. f ns = Some n \ TMC_yields_num_res tm ns n ))" unfolding TMC_yields_num_res_def by (simp add: turing_computable_total_on_def) (* Major consequences of the definition *) lemma turing_computable_total_on_UNIV_imp_turing_computable_total: "turing_computable_total_on f UNIV \ turing_computable_total f" by (simp add: turing_computable_total_on_unfolded_into_TMC_yields_condition turing_computable_total_unfolded_into_TMC_yields_condition) end diff --git a/thys/Universal_Turing_Machine/TuringComputable_aux.thy b/thys/Universal_Turing_Machine/TuringComputable_aux.thy --- a/thys/Universal_Turing_Machine/TuringComputable_aux.thy +++ b/thys/Universal_Turing_Machine/TuringComputable_aux.thy @@ -1,158 +1,158 @@ (* Title: thys/TuringComputable_aux.thy Author: Franz Regensburger (FABR) 04/2022 *) theory TuringComputable_aux imports TuringComputable begin (* Auxiliary lemmas that are not used at the moment *) lemma TMC_has_num_res_impl_result: assumes "TMC_has_num_res p ns" shows "\stp k n l. (steps0 (1, ([], )) p stp) = (0, Bk \ k, @ Bk \ l)" using Hoare_halt_on_numeral_imp_result using TMC_has_num_res_def assms by blast lemma TMC_has_num_res_impl_result_rev: assumes "\stp k n l. (steps0 (1, ([], )) p stp) = (0, Bk \ k, @ Bk \ l)" shows "TMC_has_num_res p ns" unfolding TMC_has_num_res_def using Hoare_halt_on_numeral_imp_result_rev using assms tape_of_list_def by auto corollary TMC_has_num_res_impl_result_iff: "(TMC_has_num_res p ns) \ (\stp k n l. (steps0 (1, ([], )) p stp) = (0, Bk \ k, @ Bk \ l))" using TMC_has_num_res_impl_result TMC_has_num_res_impl_result_rev by blast lemma TMC_has_num_res_impl_unique_result: assumes "TMC_has_num_res p ns" shows "\!n. \stp k l. steps0 (1,[], ) p stp = (0, Bk \ k, @ Bk \ l)" proof (rule ex_ex1I) show "\n stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" using assms TMC_has_num_res_impl_result by blast next show "\n y. \\stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l); \stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)\ \ n = y" - by (smt before_final is_final_eq le_less least_steps less_Suc_eq not_less_iff_gr_or_eq snd_conv unique_decomp_std_tap) + by (smt (verit) before_final is_final_eq le_less least_steps less_Suc_eq not_less_iff_gr_or_eq snd_conv unique_decomp_std_tap) qed lemma TMC_has_num_res_impl_unique_result_rev: assumes "\!n. \stp k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" shows "TMC_has_num_res p ns" proof - from assms have "\stp n k l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" by blast then show ?thesis using TMC_has_num_res_impl_result_rev by blast qed corollary TMC_has_num_res_impl_unique_result_iff: "TMC_has_num_res p ns \ (\!n.\stp k l. (steps0 (1, ([], )) p stp) = (0, Bk \ k, @ Bk \ l))" using TMC_has_num_res_impl_unique_result TMC_has_num_res_impl_unique_result_rev by blast subsection \Hoare halt with single numeral result: least number of steps and uniqueness\ lemma TMC_has_num_res_impl_result_least: assumes "TMC_has_num_res p ns" and "st = (LEAST m. is_final (steps0 (1, [], ) p m))" shows "\ k n l. (steps0 (1, ([], )) p st) = (0, Bk \ k, @ Bk \ l)" proof (rule exE) from assms(1) show "\stp k n l. (steps0 (1, ([], )) p stp) = (0, Bk \ k, @ Bk \ l)" by (rule TMC_has_num_res_impl_result) next fix stp assume "\k n l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" then obtain k n l where w: "steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l)" by blast then have "\ n'. (\n'' < n'. \ is_final (steps0 (1, ([], )) p n'')) \ (\n'' \ n'. is_final (steps0 (1, ([], )) p n''))" by (rule least_steps) then obtain ln where w_ln:" (\n'' < ln. \ is_final (steps0 (1, ([], )) p n'')) \ (\n'' \ ln. is_final (steps0 (1, ([], )) p n''))" by blast from w_ln have F_ln: "is_final (steps0 (1, [], ) p ln)" by auto have F_ln_eq: "(LEAST m. is_final (steps0 (1, [], ) p m)) = ln" proof (rule Least_equality) from w_ln show "is_final (steps0 (1, [], ) p ln)" by auto next fix y assume "is_final (steps0 (1, [], ) p y)" show "ln \ y" proof (rule ccontr) assume "\ ln \ y" then have "y < ln" by auto with w_ln have "\ is_final (steps0 (1, ([], )) p y)" by auto with \is_final (steps0 (1, [], ) p y)\ show False by blast qed qed with assms(2) have "st = ln" by auto with F_ln have F_st: "is_final (steps0 (1, [], ) p st)" by auto from w have F_stp: "is_final (steps0 (1, [], ) p stp)" by auto have "steps0 (1, [], ) p st = (0, Bk \ k, @ Bk \ l)" proof (cases "stp < st") case True then have "stp < st" . with F_stp and assms(2) have False using not_less_Least by auto then show ?thesis by auto next case False then have "st \ stp" by auto then show ?thesis proof (cases "steps0 (1, [], ) p st") case (fields a b c) then have "steps0 (1, [], ) p st = (a, b, c)" . moreover with F_st have S0: "a = 0" using is_final_eq[THEN iffD1] by auto ultimately have "steps0 (1, [], ) p st = (0, b, c)" by auto from this and \st \ stp\ have "steps0 (1, [], ) p stp = (0, b, c)" by (rule stable_config_after_final_ge') with w have "b = Bk \ k \ c= @ Bk \ l" by auto with S0 and \steps0 (1, [], ) p st = (a, b, c)\ show ?thesis by auto qed qed then show "\k n l. steps0 (1, [], ) p st = (0, Bk \ k, @ Bk \ l)" by auto qed lemma TMC_has_num_res_impl_result_is_unique: assumes "TMC_has_num_res p ns" and "st = (LEAST m. is_final (steps0 (1,[], ) p m))" shows "\!n. \k l. steps0 (1, [], ) p st = (0, Bk \ k, @ Bk \ l)" proof (rule ex_ex1I) show "\n k l. steps0 (1, [], ) p st = (0, Bk \ k, @ Bk \ l)" proof - from assms have "\ k n l. (steps0 (1, ([], )) p st) = (0, Bk \ k, @ Bk \ l)" by (rule TMC_has_num_res_impl_result_least) then obtain k n l where w_k_n_l: "(steps0 (1, ([], )) p st) = (0, Bk \ k, @ Bk \ l)" by blast then show "\n k l. steps0 (1, [], ) p st = (0, Bk \ k, @ Bk \ l)" by auto qed next fix n::nat fix y::nat assume A1: "\k l. steps0 (1, [], ) p st = (0, Bk \ k, @ Bk \ l)" and A2: "\k l. steps0 (1, [], ) p st = (0, Bk \ k, @ Bk \ l)" from A1 obtain k1 l1 where w_k1_l1: "steps0 (1, [], ) p st = (0, Bk \ k1, @ Bk \ l1)" by blast from A2 obtain k2 l2 where w_k2_l2: "steps0 (1, [], ) p st = (0, Bk \ k2, @ Bk \ l2)" by blast from w_k1_l1 and w_k2_l2 have "(0, Bk \ k1, @ Bk \ l1) = (0, Bk \ k2, @ Bk \ l2)" by auto then have "(Bk \ k1, @ Bk \ l1) = (Bk \ k2, @ Bk \ l2)" by auto then have "k1=k2 \ n=y \ l1=l2" by (rule unique_decomp_std_tap) then show "n = y" by auto qed end diff --git a/thys/Universal_Turing_Machine/Turing_HaltingConditions.thy b/thys/Universal_Turing_Machine/Turing_HaltingConditions.thy --- a/thys/Universal_Turing_Machine/Turing_HaltingConditions.thy +++ b/thys/Universal_Turing_Machine/Turing_HaltingConditions.thy @@ -1,594 +1,594 @@ (* Title: thys/Turing_HaltingConditions.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten Further contributions by Franz Regensburger (FABR) 02/2022 *) section \Halting Conditions and Standard Halting Configuration\ theory Turing_HaltingConditions imports Turing_Hoare begin subsection \Looping of Turing Machines\ definition TMC_loops :: "tprog0 \ nat list \ bool" where "TMC_loops p ns \ (\stp.\ is_final (steps0 (1, [], ) p stp))" subsection \Reaching the Final State\ definition reaches_final :: "tprog0 \ nat list \ bool" where "reaches_final p ns \ \(\tap. tap = ([], ))\ p \(\tap. True)\" text \The definition @{term reaches_final} and all lemmas about it are only needed for the special halting problem K0. \ lemma True_holds_for_all: "(\tap. True) holds_for c" by (cases c)(auto) lemma reaches_final_iff: "reaches_final p ns \ (\n. is_final (steps0 (1, ([], )) p n))" by (auto simp add: reaches_final_def Hoare_halt_def True_holds_for_all) text \Some lemmas about reaching the Final State.\ lemma Hoare_halt_from_init_imp_reaches_final: assumes "\\tap. tap = ([], )\ p \Q\" shows "reaches_final p ns" proof - from assms have "\tap. tap = ([], ) \ (\n. is_final (steps0 (1, tap) p n))" using Hoare_halt_def by auto then show ?thesis using reaches_final_iff by auto qed lemma Hoare_unhalt_impl_not_reaches_final: assumes "\(\tap. tap = ([], ))\ p \" shows "\(reaches_final p ns)" proof assume "reaches_final p ns" then have "(\n. is_final (steps0 (1, ([], )) p n))" by (auto simp add: reaches_final_iff) then obtain n where w_n: "is_final (steps0 (1, ([], )) p n)" by blast from assms have "\tap. (\tap. tap = ([], )) tap \ (\ n . \ (is_final (steps0 (1, tap) p n)))" by (auto simp add: Hoare_unhalt_def) then have "\ (is_final (steps0 (1, ([], )) p n))" by blast with w_n show False by auto qed subsection \What is a Standard Tape\ text \A tape is called {\em standard}, if the left tape is empty or contains only blanks and the right tape contains a single nonempty block of strokes (occupied cells) followed by zero or more blanks.. Thus, by definition of left and right tape, the head of the machine is always scanning the first cell of this single block of strokes. We extend the notion of a standard tape to lists of numerals as well. \ definition std_tap :: "tape \ bool" where "std_tap tap \ (\k n l. tap = (Bk \ k, @ Bk \ l))" definition std_tap_list :: "tape \ bool" where "std_tap_list tap \ (\k ml l. tap = (Bk \ k, @ Bk \ l))" lemma "std_tap tap \ std_tap_list tap" unfolding std_tap_def std_tap_list_def by (metis tape_of_list_def tape_of_nat_list.simps(2)) text \A configuration @{term "(st, l, r)"} of a Turing machine is called a {\em standard configuration}, if the state @{term "st"} is the final state $0$ and the @{term "(l, r)"} is a standard tape. \ (* A duplicate from UTM just for comparison of concepts *) definition TSTD':: "config \ bool" where "TSTD' c = ((let (st, l, r) = c in st = 0 \ (\ m. l = Bk\(m)) \ (\ rs n. r = Oc\(Suc rs) @ Bk\(n))))" (* Relate definition of TSTD' to std_tap *) lemma "TSTD' (st, l, r) = ((st = 0) \ std_tap (l,r))" unfolding TSTD'_def std_tap_def proof - have "(let (st, l, r) = (st, l, r) in st = 0 \ (\m. l = Bk \ m) \ (\rs n. r = Oc \ Suc rs @ Bk \ n)) = (st = 0 \ (\m. l = Bk \ m) \ (\rs n. r = Oc \ Suc rs @ Bk \ n))" by auto also have "... = (st = 0 \ (\m. l = Bk \ m) \ (\n la. r = ( @ Bk \ la)))" by (auto simp add: tape_of_nat_def) finally have "(let (st, l, r) = (st, l, r) in st = 0 \ (\m. l = Bk \ m) \ (\rs n. r = Oc \ Suc rs @ Bk \ n)) = (st = 0 \ (\m. l = Bk \ m) \ (\n la. r = ( @ Bk \ la)))" by (auto simp add: tape_of_nat_def) moreover have "((\m. l = Bk \ m) \ (\n la. r = @ Bk \ la)) = (\k n la. (l, r) = (Bk \ k, @ Bk \ la))" by auto ultimately show "(let (st, l, r) = (st, l, r) in st = 0 \ (\m. l = Bk \ m) \ (\rs n. r = Oc \ Suc rs @ Bk \ n)) = (st = 0 \ (\k n la. (l, r) = (Bk \ k, @ Bk \ la)))" by blast qed subsection \What does Hoare\_halt mean in general?\ text \We say {\em in general} because the result computed on the right tape is not necessarily a numeral but some arbitrary component @{term "r'"} . \ lemma Hoare_halt2_iff: "\\tap. \kl ll. tap = (Bk \ kl, r @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\ \ (\kl ll. \n. is_final (steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n) \ (\kr lr. steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n = (0, Bk \ kr, r' @ Bk \ lr)))" proof assume "\\tap. \kl ll. tap = (Bk \ kl, r @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" then show "\kl ll. \n. is_final (steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n) \ (\kr lr. steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n = (0, Bk \ kr, r' @ Bk \ lr))" - by (smt Hoare_halt_def Pair_inject holds_for.elims(2) is_final.elims(2)) + using Hoare_halt_E0 is_finalI by force next assume "\kl ll. \n. is_final (steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n) \ (\kr lr. steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n = (0, Bk \ kr, r' @ Bk \ lr))" then show "\\tap. \kl ll. tap = (Bk \ kl, r @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" unfolding Hoare_halt_def using holds_for.simps by fastforce qed lemma Hoare_halt_D: assumes "\\tap. \kl ll. tap = (Bk \ kl, r @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" shows "\n. is_final (steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n) \ (\kr lr. steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n = (0, Bk \ kr, r' @ Bk \ lr))" proof - from assms show "\n. is_final (steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n) \ (\kr lr. steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n = (0, Bk \ kr, r' @ Bk \ lr))" by (simp add: Hoare_halt2_iff is_final_eq) qed lemma Hoare_halt_I2: assumes "\kl ll. \n. is_final (steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n) \ (\kr lr. steps0 (1, (Bk \ kl, r @ Bk \ ll)) p n = (0, Bk \ kr, r' @ Bk \ lr))" shows "\\tap. \kl ll. tap = (Bk \ kl, r @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" unfolding Hoare_halt_def using assms holds_for.simps by fastforce lemma Hoare_halt_D_Nil: assumes "\\tap. tap = ([], r)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" shows "\n. is_final (steps0 (1, ([], r)) p n) \ (\kr lr. steps0 (1, ([], r)) p n = (0, Bk \ kr, r' @ Bk \ lr))" proof - from assms have "\\tap. tap = (Bk \ 0, r @ Bk \ 0)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" by simp then have "\n. is_final (steps0 (1, (Bk \ 0, r @ Bk \ 0)) p n) \ (\kr lr. steps0 (1, (Bk \ 0, r @ Bk \ 0)) p n = (0, Bk \ kr, r' @ Bk \ lr))" using Hoare_halt_E0 append_self_conv assms is_final_eq old.prod.inject prod.inject replicate_0 by force then show ?thesis by auto qed lemma Hoare_halt_I2_Nil: assumes "\n. is_final (steps0 (1, ([], r )) p n) \ (\kr lr. steps0 (1, ([], r )) p n = (0, Bk \ kr, r' @ Bk \ lr))" shows "\\tap. tap = ([], r)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" proof - from assms have "\n. is_final (steps0 (1, (Bk \ 0, r @ Bk \ 0)) p n) \ (\kr lr. steps0 (1, (Bk \ 0, r @ Bk \ 0)) p n = (0, Bk \ kr, r' @ Bk \ lr))" by auto then have "\\tap. tap = (Bk \ 0, r @ Bk \ 0)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\" using Hoare_halt_iff by auto then show ?thesis by auto qed lemma Hoare_halt2_Nil_iff: "\\tap. tap = ([], r)\ p \\tap. \kr lr. tap = (Bk \ kr, r' @ Bk \ lr)\ \ (\n. is_final (steps0 (1, ([], r)) p n) \ (\kr lr. steps0 (1, ([], r)) p n = (0, Bk \ kr, r' @ Bk \ lr)))" using Hoare_halt_D_Nil Hoare_halt_I2_Nil by blast corollary Hoare_halt_left_tape_Nil_imp_All_left_and_right: assumes "\\tap. tap = ([] , r )\ p \\tap. \k l. tap = (Bk \ k , r' @ Bk \ l)\" shows "\\tap. \x y. tap = (Bk \ x, r @ Bk \ y)\ p \\tap. \k l. tap = (Bk \ k , r' @ Bk \ l)\" proof - from assms have "\n. is_final (steps0 (1, ([], r)) p n) \ (\k l. steps0 (1, ([], r)) p n = (0, Bk \ k, r' @ Bk \ l))" using Hoare_halt_D_Nil by auto then have "\x y. \n. is_final (steps0 (1, (Bk \ x, r @ Bk \ y)) p n) \ (\k l. steps0 (1, (Bk \ x, r @ Bk \ y)) p n = (0, Bk \ k, r' @ Bk \ l))" using ex_steps_left_tape_Nil_imp_All_left_and_right using is_final.simps by force then show ?thesis using Hoare_halt_I2 by auto qed subsubsection \What does Hoare\_halt with a numeral list result mean?\ text \About computations that result in numeral lists on the final right tape.\ (* Adding trailing blanks to the initial right tape needs tough lemmas (see BlanksDoNotMatter) *) lemma TMC_has_num_res_list_without_initial_Bks_imp_TMC_has_num_res_list_after_adding_Bks_to_initial_right_tape: " \\tap. tap = ([], )\ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\ \ \\tap. \ll. tap = ([], @ Bk \ ll)\ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\" proof - assume A: "\\tap. tap = ([], )\ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\" then have "\n. is_final (steps0 (1, ([], )) p n) \ (\ms kr lr. steps0 (1, ([], )) p n = (0, Bk \ kr, @ Bk \ lr))" using Hoare_halt_E0 is_finalI by force then obtain stp where w_stp: "is_final (steps0 (1, ([], )) p stp) \ (\ms kr lr. steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr))" by blast then obtain ms where "\kr lr. steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr)" by blast then have "\ll. \kr lr. steps0 (1, ([], @ Bk \ ll)) p stp = (0, Bk \ kr, @ Bk \ lr)" using ex_steps_left_tape_Nil_imp_All_left_and_right steps_left_tape_ShrinkBkCtx_to_NIL by blast then have "\ll. is_final (steps0 (1, ([], @ Bk \ ll)) p stp) \ (\ms kr lr. steps0 (1, ([], @ Bk \ ll)) p stp = (0, Bk \ kr, @ Bk \ lr))" by (metis is_finalI) then have "\tap. (\ll. tap = ([], @ Bk \ ll)) \ (\n. is_final (steps0 (1, tap) p n) \ (\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)) holds_for steps0 (1, tap) p n)" using holds_for.simps by force then show ?thesis unfolding Hoare_halt_def by auto qed (* Removing trailing blanks on the initial right tape (is easy) *) lemma TMC_has_num_res_list_without_initial_Bks_iff_TMC_has_num_res_list_after_adding_Bks_to_initial_right_tape: "\\tap. tap = ([], ) \ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\ \ \\tap. \ll. tap = ([], @ Bk \ ll)\ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\" proof assume "\\tap. tap = ([], ) \ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\" then show "\\tap. \ll. tap = ([], @ Bk \ ll)\ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\" using TMC_has_num_res_list_without_initial_Bks_imp_TMC_has_num_res_list_after_adding_Bks_to_initial_right_tape by blast next assume "\\tap. \ll. tap = ([], @ Bk \ ll)\ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\" then show "\\tap. tap = ([], ) \ p \\tap. \ms kr lr. tap = (Bk \ kr, @ Bk \ lr)\" by (simp add: Hoare_halt_def assert_imp_def) qed (* In addition we may add or remove blanks on the initial left tape *) lemma TMC_has_num_res_list_without_initial_Bks_imp_TMC_has_num_res_list_after_adding_Bks_to_initial_left_and_right_tape: " \\tap. tap = ([], )\ p \\tap.\kr lr. tap = (Bk \ kr, @ Bk \ lr)\ \ \\tap. \kl ll. tap = (Bk \ kl, @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, @ Bk \ lr)\" using Hoare_halt_left_tape_Nil_imp_All_left_and_right by auto (* proof - assume A: "\\tap. tap = ([], )\ p \\tap.\kr lr. tap = (Bk \ kr, @ Bk \ lr)\" then have "\stp. is_final (steps0 (1, ([], )) p stp) \ (\ kr lr. steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr))" using Hoare_halt2_Nil_iff tape_of_list_def tape_of_nat_list.simps(2) by blast then obtain stp where w_stp: "is_final (steps0 (1, ([], )) p stp) \ (\ kr lr. steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr))" by blast then have "\kl ll.\kr lr. steps0 (1, Bk \ kl, @ Bk \ ll) p stp = (0, Bk \ kr, @ Bk \ lr)" using ex_steps_left_tape_Nil_imp_All_left_and_right steps_left_tape_ShrinkBkCtx_to_NIL by blast then have "\kl ll. is_final (steps0 (1, Bk \ kl, @ Bk \ ll) p stp) \ (\kr lr. steps0 (1, Bk \ kl , @ Bk \ ll) p stp = (0, Bk \ kr, @ Bk \ lr))" by (metis is_finalI) then have "\tap. (\kl ll. tap = (Bk \ kl, @ Bk \ ll) ) \ (\stp. is_final (steps0 (1, tap) p stp) \ (\tap. \kr lr. tap = (Bk \ kr, @ Bk \ lr)) holds_for steps0 (1, tap) p stp)" using holds_for.simps by force then show ?thesis unfolding Hoare_halt_def by auto qed *) lemma TMC_has_num_res_list_without_initial_Bks_iff_TMC_has_num_res_list_after_adding_Bks_to_initial_left_and_right_tape: " \\tap. tap = ([], )\ p \\tap.\kr lr. tap = (Bk \ kr, @ Bk \ lr)\ \ \\tap. \kl ll. tap = (Bk \ kl, @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, @ Bk \ lr)\" proof assume "\\tap. tap = ([], )\ p \\tap.\kr lr. tap = (Bk \ kr, @ Bk \ lr)\" then show "\\tap. \kl ll. tap = (Bk \ kl, @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, @ Bk \ lr)\" using TMC_has_num_res_list_without_initial_Bks_imp_TMC_has_num_res_list_after_adding_Bks_to_initial_left_and_right_tape by auto next assume "\\tap. \kl ll. tap = (Bk \ kl, @ Bk \ ll)\ p \\tap. \kr lr. tap = (Bk \ kr, @ Bk \ lr)\" then show "\\tap. tap = ([], )\ p \\tap.\kr lr. tap = (Bk \ kr, @ Bk \ lr)\" by (simp add: Hoare_halt_def assert_imp_def) qed subsection \Halting in a Standard Configuration\ subsubsection \Definition of Halting in a Standard Configuration\ text \The predicates \TMC_has_num_res p ns\ and \TMC_has_num_list_res\ describe that a run of the Turing program \p\ on input \ns\ reaches the final state 0 and the final tape produced thereby is standard. Thus, the computation of the Turing machine \p\ produced a result, which is either a single numeral or a list of numerals. Since trailing blanks on the initial left or right tape do not matter, we may restrict our definitions to the case where the initial left tape is empty and there are no trailing blanks on the initial right tape! \ definition TMC_has_num_res :: "tprog0 \ nat list \ bool" where "TMC_has_num_res p ns \ \ \tap. tap = ([], ) \ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" lemma TMC_has_num_res_iff: "TMC_has_num_res p ns \ (\stp. is_final (steps0 (1, [],) p stp) \ (\k n l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk \ l)))" unfolding TMC_has_num_res_def proof assume "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" then show "\stp. is_final (steps0 (1, [], ) p stp) \ (\k n l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l))" - by (smt Hoare_halt_def holds_for.elims(2) is_final.elims(2) is_final_eq) + by (smt (verit, best) Hoare_halt_E0 Hoare_halt_impl_not_Hoare_unhalt Hoare_unhalt_def is_finalI tape_of_list_def tape_of_nat_def) next assume "\stp. is_final (steps0 (1, [], ) p stp) \ (\k n l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l))" then show "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" by (metis (mono_tags, lifting) Hoare_halt_def holds_for.simps) qed (* for clarification lemma TMC_has_num_res_iff': "TMC_has_num_res p ns \ (\stp k n l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk \ l))" by (smt TMC_has_num_res_iff is_finalI steps_0 steps_add) *) (* --- *) definition TMC_has_num_list_res :: "tprog0 \ nat list \ bool" where "TMC_has_num_list_res p ns \ \\tap. tap = ([], )\ p \\tap. \kr ms lr. tap = (Bk \ kr, @ Bk \ lr)\" lemma TMC_has_num_list_res_iff: "TMC_has_num_list_res p ns \ (\stp. is_final (steps0 (1, [],) p stp) \ (\k ms l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk \ l)))" unfolding TMC_has_num_list_res_def proof assume "\\tap. tap = ([], )\ p \\tap. \k ms l. tap = (Bk \ k, @ Bk \ l)\" then show "\stp. is_final (steps0 (1, [], ) p stp) \ (\k ms l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l))" - by (smt Hoare_halt_def holds_for.elims(2) is_final.elims(2) is_final_eq) + using Hoare_halt_E0 is_finalI by force next assume "\stp. is_final (steps0 (1, [], ) p stp) \ (\k ms l. steps0 (1, [], ) p stp = (0, Bk \ k, @ Bk \ l))" then show "\\tap. tap = ([], )\ p \\tap. \k ms l. tap = (Bk \ k, @ Bk \ l)\" by (metis (mono_tags, lifting) Hoare_halt_def holds_for.simps) qed (* for clarification lemma TMC_has_num_list_res_iff': "TMC_has_num_list_res p ns \ (\stp k ms l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk \ l))" by (smt TMC_has_num_list_res_iff is_finalI steps_0 steps_add) *) subsubsection \Relation between TMC\_has\_num\_res and TMC\_has\_num\_list\_res\ text \A computation of a Turing machine, which started on a list of numerals and halts in a standard configuration with a single numeral result is a special case of a halt in a standard configuration that halts with a list of numerals.\ (* FABR: this is an important lemma, since it motivates and validates our extension * of various concepts to lists of natural numbers or numerals *) theorem TMC_has_num_res_imp_TMC_has_num_list_res: "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\ \ \\tap. tap = ([], )\ p \\tap. \kr ms lr. tap = (Bk \ kr, @ Bk \ lr)\" proof - assume A: "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" then have "\stp. is_final (steps0 (1, ([], )) p stp) \ (\n kr lr. steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr))" - by (smt Hoare_halt_def holds_for.elims(2) is_final.elims(2) is_final_eq) + using Hoare_halt_E0 is_finalI by force then obtain stp where w_stp: "is_final (steps0 (1, ([], )) p stp) \ (\n kr lr. steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr))" by blast then have "(\n kr lr. steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr))" by auto then obtain n kr lr where "steps0 (1, ([], )) p stp = (0, Bk \ kr, @ Bk \ lr)" by blast then have "steps0 (1, ([], )) p stp = (0, Bk \ kr, <[n::nat]> @ Bk \ lr)" by (simp add: tape_of_list_def) then have "is_final (steps0 (1, ([], )) p stp) \ (\kr lr. (steps0 (1, ([], )) p stp) = (0, Bk \ kr, <[n::nat]> @ Bk \ lr))" by (simp add: is_final_eq) then have "is_final (steps0 (1, ([], )) p stp) \ (\ms kr lr. (steps0 (1, ([], )) p stp) = (0, Bk \ kr, @ Bk \ lr))" by blast then show "\\tap. tap = ([], )\ p \\tap. \kr ms lr. tap = (Bk \ kr, @ Bk \ lr)\" by (metis (mono_tags, lifting) Hoare_halt_def holds_for.simps) qed corollary TMC_has_num_res_imp_TMC_has_num_list_res': "TMC_has_num_res p ns \ TMC_has_num_list_res p ns" unfolding TMC_has_num_res_def TMC_has_num_list_res_def using TMC_has_num_res_imp_TMC_has_num_list_res by auto subsubsection \Convenience Lemmas for Halting Problems \ lemma Hoare_halt_with_Oc_imp_std_tap: assumes "\(\tap. tap = ([], ))\ p \\tap. \k l. tap = (Bk \ k, [Oc] @ Bk \ l)\" shows "TMC_has_num_res p ns" unfolding TMC_has_num_res_def proof - from assms have F0: "\(\tap. tap = ([], ))\ p \\tap. \k l. tap = (Bk \ k, <0::nat> @ Bk \ l)\" by (auto simp add: tape_of_nat_def) show "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" - by (smt F0 Hoare_halt_E0 Hoare_halt_I0) + using F0 Hoare_haltI Hoare_halt_E0 Hoare_halt_iff by fastforce qed lemma Hoare_halt_with_OcOc_imp_std_tap: assumes "\(\tap. tap = ([], ))\ p \\tap. \k l. tap = (Bk \ k, [Oc, Oc] @ Bk \ l)\" shows "TMC_has_num_res p ns" unfolding TMC_has_num_res_def proof - from assms have "\(\tap. tap = ([], ))\ p \\tap. \k l. tap = (Bk \ k, <1::nat> @ Bk \ l)\" by (auto simp add: tape_of_nat_def) then show "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" - by (smt Hoare_haltE Hoare_haltI holds_for.elims(2) holds_for.simps) + using Hoare_halt_E0 Hoare_halt_iff by fastforce qed subsubsection \Hoare\_halt on numeral lists with single numeral result\ (* For automation we use Hoare triples directly instead of the concept TMC_has_num_res. * Reason: reduce number of theorems used by the simplifier. *) lemma Hoare_halt_on_numeral_imp_result: (* special case of Hoare_halt_D *) assumes "\(\tap. tap = ([], ))\ p \(\tap. (\k n l. tap = (Bk \ k, @ Bk \ l)))\" shows "\stp k n l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l)" using TMC_has_num_res_def TMC_has_num_res_iff assms by blast lemma Hoare_halt_on_numeral_imp_result_rev: (* special case of Hoare_halt_I2 *) assumes "\stp k n l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l)" shows "\(\tap. tap = ([], ))\ p \(\tap. (\k n l. tap = (Bk \ k, @ Bk \ l)))\" using TMC_has_num_res_def TMC_has_num_res_iff assms is_final_eq by force lemma Hoare_halt_on_numeral_imp_result_iff: (* special case of Hoare_halt2_iff *) "\(\tap. tap = ([], ))\ p \(\tap. (\k n l. tap = (Bk \ k, @ Bk \ l)))\ \ (\stp k n l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l))" using Hoare_halt_on_numeral_imp_result Hoare_halt_on_numeral_imp_result_rev by blast subsubsection \Hoare\_halt on numeral lists with numeral list result\ (* For automation we use Hoare triples directly instead of the concept TMC_has_num_list_res. * Reason: reduce number of theorems used by the simplifier. *) lemma Hoare_halt_on_numeral_imp_list_result: (* special case of Hoare_halt_D *) assumes "\(\tap. tap = ([], ))\ p \(\tap. (\k ms l. tap = (Bk \ k, @ Bk \ l)))\" shows "\stp k ms l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l)" using TMC_has_num_list_res_def TMC_has_num_list_res_iff assms by blast lemma Hoare_halt_on_numeral_imp_list_result_rev: (* special case of Hoare_halt_I2 *) assumes "\stp k ms l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l)" shows "\(\tap. tap = ([], ))\ p \(\tap. (\k ms l. tap = (Bk \ k, @ Bk \ l)))\" by (metis (mono_tags, lifting) Hoare_haltI assms holds_for.simps is_finalI) lemma Hoare_halt_on_numeral_imp_list_result_iff: (* special case of Hoare_halt2_iff *) "\(\tap. tap = ([], ))\ p \(\tap. (\k ms l. tap = (Bk \ k, @ Bk \ l)))\ \ (\stp k ms l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l))" using Hoare_halt_on_numeral_imp_list_result Hoare_halt_on_numeral_imp_list_result_rev by blast subsection \Trailing left blanks do not matter for computations with result\ (* adding trailing blanks on the initial left tape needs tough lemmas (see BlanksDoNotMatter) *) lemma TMC_has_num_res_NIL_impl_TMC_has_num_res_with_left_BKs: "\(\tap. tap = ([], ))\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \ \ \(\tap. \z. tap = (Bk\z, ))\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" proof - assume "\\tap. tap = ([], )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" then have "\stp k n l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l)" by (rule Hoare_halt_on_numeral_imp_result) then obtain n where "\stp k l. steps0 (1, [],) p stp = (0, Bk \ k, @ Bk\ l)" by blast then have "\z. \stp k l. (steps0 (1, (Bk\z, )) p stp) = (0, Bk \ k, @ Bk \ l)" by (rule ex_steps_left_tape_Nil_imp_All) then have "\(\tap. \z. tap = (Bk\z, ))\ p \ (\tap. (\k l. tap = (Bk \ k, @ Bk \ l))) \" by (rule Hoare_halt_add_Bks_left_tape_L2) then show "\\tap. \z. tap = (Bk \ z, )\ p \\tap. \k n l. tap = (Bk \ k, @ Bk \ l)\" using Hoare_halt_iff \\z. \stp k l. steps0 (1, Bk \ z, ) p stp = (0, Bk \ k, @ Bk \ l)\ by fastforce qed (* removing trailing blanks on the initial left tape (is easy) *) corollary TMC_has_num_res_NIL_iff_TMC_has_num_res_with_left_BKs: " \(\tap. tap = ([], ))\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \ \ \(\tap. \z. tap = (Bk\z, ))\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" proof assume "\\tap. tap = ([], )\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" then show "\\tap. \z. tap = (Bk \ z, )\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" using TMC_has_num_res_NIL_impl_TMC_has_num_res_with_left_BKs by blast next assume "\\tap. \z. tap = (Bk \ z, )\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" then show "\\tap. tap = ([], )\ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" by (simp add: Hoare_halt_def assert_imp_def) qed subsection \About Turing Computations and the result they yield\ definition TMC_yields_num_res :: "tprog0 \ nat list \ nat \ bool" where "TMC_yields_num_res tm ns n \ (\stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l))" definition TMC_yields_num_list_res :: "tprog0 \ nat list \ nat list \ bool" where "TMC_yields_num_list_res tm ns ms \ (\stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l))" (* This is for documentation and explanation: TMC_yields_num_res_unfolded_into_Hoare_halt *) lemma TMC_yields_num_res_unfolded_into_Hoare_halt: "TMC_yields_num_res tm ns n \ \(\tap. tap = ([], ))\ tm \(\tap. \k l. tap = (Bk \ k, @ Bk\ l))\" - by (smt Hoare_halt_iff TMC_yields_num_res_def) + by (smt (verit, ccfv_threshold) Hoare_halt_iff TMC_yields_num_res_def) (* This is for documentation and explanation: TMC_yields_num_list_res_unfolded_into_Hoare_halt *) lemma TMC_yields_num_list_res_unfolded_into_Hoare_halt: "TMC_yields_num_list_res tm ns ms \ \(\tap. tap = ([], ))\ tm \(\tap. \k l. tap = (Bk \ k, @ Bk\ l))\" - by (smt Hoare_halt_iff TMC_yields_num_list_res_def) + by (smt (verit, ccfv_threshold) Hoare_halt_E0 Hoare_halt_def Hoare_halt_iff TMC_yields_num_list_res_def) (* A variant of rule Hoare_plus_halt using TMC_yields_num_list_res and TMC_yields_num_res *) lemma TMC_yields_num_res_Hoare_plus_halt: assumes "TMC_yields_num_list_res tm1 nl r1" and "TMC_yields_num_res tm2 r1 r2" and "composable_tm0 tm1" shows "TMC_yields_num_res (tm1 |+| tm2) nl r2" proof - from assms have "\\tap. tap = ([], )\ tm1 \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" using TMC_yields_num_list_res_unfolded_into_Hoare_halt by auto moreover from assms have "\\tap. tap = ([], )\ tm2 \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" using TMC_yields_num_res_unfolded_into_Hoare_halt Hoare_halt_def Hoare_halt_iff TMC_yields_num_res_def by blast ultimately have "\\tap. tap = ([], )\ (tm1 |+| tm2) \\tap. \k l. tap = (Bk \ k, @ Bk \ l)\" using \composable_tm0 tm1\ using Hoare_halt_left_tape_Nil_imp_All_left_and_right Hoare_plus_halt by (simp add: tape_of_list_def) then show ?thesis using TMC_yields_num_res_unfolded_into_Hoare_halt by auto qed (* TODO test mixfix notation for TMC_yields_num_res: \tm: ns \ n\ TODO test mixfix notation for TMC_has_num_res: \tm: ns\\ Trial for notation of TMC_yields_num_res tm ns n: "TMC_yields_num_res tm ns n \ (\stp k l. (steps0 (1, ([], )) tm stp) = (0, Bk \ k, @ Bk \ l))" ns \\ n \tm: ns \ n\ that's good \tm: ns \ n\ that's good as well TMC_has_num_res "TMC_has_num_res p ns \ \ \tap. tap = ([], ) \ p \ \tap. (\k n l. tap = (Bk \ k, @ Bk \ l)) \" \tm: ns\\ *) end diff --git a/thys/Universal_Turing_Machine/UTM.thy b/thys/Universal_Turing_Machine/UTM.thy --- a/thys/Universal_Turing_Machine/UTM.thy +++ b/thys/Universal_Turing_Machine/UTM.thy @@ -1,4749 +1,4749 @@ (* Title: thys/UTM.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten Modifications: Franz Regensburger (FABR) 08/2022 - added some auxiliary lemmas - added subsections *) chapter \Construction of a Universal Turing Machine\ theory UTM imports Recursive Abacus UF HOL.GCD Turing_Hoare begin (* Initialize simpset: the following proofs depend on a cleaned up simpset *) (* declare adjust.simps[simp del] declare seq_tm.simps [simp del] declare shift.simps[simp del] declare composable_tm.simps[simp del] declare step.simps[simp del] declare steps.simps[simp del] declare fetch.simps[simp del] *) section \Wang coding of input arguments\ text \ The direct compilation of the universal function \rec_F\ can not give us the \utm\, because \rec_F\ is of arity 2, where the first argument represents the Gödel coding of the TM being simulated and the second argument represents the right number (in Wang's coding) of the TM tape. (Notice, the left number is always \0\ at the very beginning). However, the \utm\ needs to simulate the execution of any TM which may take many input arguments. Therefore, an initialization TM needs to run before the TM compiled from \rec_F\, and the sequential composition of these two TMs will give rise to the \utm\ we are seeking. The purpose of this initialization TM is to transform the multiple input arguments of the TM being simulated into Wang's coding, so that it can be consumed by the TM compiled from \rec_F\ as the second argument. However, this initialization TM (named \wcode_tm\) can not be constructed by compiling from any recursive function, because every recursive function takes a fixed number of input arguments, while \wcode_tm\ needs to take varying number of arguments and tranform them into Wang's coding. Therefore, this section gives a direct construction of \wcode_tm\ with just some parts being obtained from recursive functions. \newlength{\basewidth} \settowidth{\basewidth}{xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx} \newlength{\baseheight} \settoheight{\baseheight}{$B:R$} \newcommand{\vsep}{5\baseheight} The TM used to generate the Wang's code of input arguments is divided into three TMs executed sequentially, namely $prepare$, $mainwork$ and $adjust$. According to the convention, the start state of ever TM is fixed to state $1$ while the final state is fixed to $0$. The input and output of $prepare$ are illustrated respectively by Figure \ref{prepare_input} and \ref{prepare_output}. \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} [tbox/.style = {draw, thick, inner sep = 5pt}] \node (0) {}; \node (1) [tbox, text height = 3.5pt, right = -0.9pt of 0] {$m$}; \node (2) [tbox, right = -0.9pt of 1] {$0$}; \node (3) [tbox, right = -0.9pt of 2] {$a_1$}; \node (4) [tbox, right = -0.9pt of 3] {$0$}; \node (5) [tbox, right = -0.9pt of 4] {$a_2$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [tbox, right = -0.9pt of 6] {$a_n$}; \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1); \end{tikzpicture}} \caption{The input of TM $prepare$} \label{prepare_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.5pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_n$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$0$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$1$}; \draw [->, >=latex, thick] (10)+(0, -4\baseheight) -- (10); \end{tikzpicture}} \caption{The output of TM $prepare$} \label{prepare_output} \end{figure} As shown in Figure \ref{prepare_input}, the input of $prepare$ is the same as the the input of \utm\, where $m$ is the Gödel coding of the TM being interpreted and $a_1$ through $a_n$ are the $n$ input arguments of the TM under interpretation. The purpose of $purpose$ is to transform this initial tape layout to the one shown in Figure \ref{prepare_output}, which is convenient for the generation of Wang's coding of $a_1, \ldots, a_n$. The coding procedure starts from $a_n$ and ends after $a_1$ is encoded. The coding result is stored in an accumulator at the end of the tape (initially represented by the $1$ two blanks right to $a_n$ in Figure \ref{prepare_output}). In Figure \ref{prepare_output}, arguments $a_1, \ldots, a_n$ are separated by two blanks on both ends with the rest so that movement conditions can be implemented conveniently in subsequent TMs, because, by convention, two consecutive blanks are usually used to signal the end or start of a large chunk of data. The diagram of $prepare$ is given in Figure \ref{prepare_diag}. \begin{figure}[h!] \centering \scalebox{0.9}{ \begin{tikzpicture} \node[circle,draw] (1) {$1$}; \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$}; \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$}; \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$}; \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$}; \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$}; \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$}; \node[circle,draw] (8) at ($(7)+(0.3\basewidth, 0)$) {$0$}; \draw [->, >=latex] (1) edge [loop above] node[above] {$S_1:L$} (1) ; \draw [->, >=latex] (1) -- node[above] {$S_0:S_1$} (2) ; \draw [->, >=latex] (2) edge [loop above] node[above] {$S_1:R$} (2) ; \draw [->, >=latex] (2) -- node[above] {$S_0:L$} (3) ; \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3) ; \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) -- node[above] {$S_0:R$} (5) ; \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5) ; \draw [->, >=latex] (5) -- node[above] {$S_0:R$} (6) ; \draw [->, >=latex] (6) edge[bend left = 50] node[below] {$S_1:R$} (5) ; \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (7) ; \draw [->, >=latex] (7) edge[loop above] node[above] {$S_0:S_1$} (7) ; \draw [->, >=latex] (7) -- node[above] {$S_1:L$} (8) ; \end{tikzpicture}} \caption{The diagram of TM $prepare$} \label{prepare_diag} \end{figure} The purpose of TM $mainwork$ is to compute Wang's encoding of $a_1, \ldots, a_n$. Every bit of $a_1, \ldots, a_n$, including the separating bits, is processed from left to right. In order to detect the termination condition when the left most bit of $a_1$ is reached, TM $mainwork$ needs to look ahead and consider three different situations at the start of every iteration: \begin{enumerate} \item The TM configuration for the first situation is shown in Figure \ref{mainwork_case_one_input}, where the accumulator is stored in $r$, both of the next two bits to be encoded are $1$. The configuration at the end of the iteration is shown in Figure \ref{mainwork_case_one_output}, where the first 1-bit has been encoded and cleared. Notice that the accumulator has been changed to $(r+1) \times 2$ to reflect the encoded bit. \item The TM configuration for the second situation is shown in Figure \ref{mainwork_case_two_input}, where the accumulator is stored in $r$, the next two bits to be encoded are $1$ and $0$. After the first $1$-bit was encoded and cleared, the second $0$-bit is difficult to detect and process. To solve this problem, these two consecutive bits are encoded in one iteration. In this situation, only the first $1$-bit needs to be cleared since the second one is cleared by definition. The configuration at the end of the iteration is shown in Figure \ref{mainwork_case_two_output}. Notice that the accumulator has been changed to $(r+1) \times 4$ to reflect the two encoded bits. \item The third situation corresponds to the case when the last bit of $a_1$ is reached. The TM configurations at the start and end of the iteration are shown in Figure \ref{mainwork_case_three_input} and \ref{mainwork_case_three_output} respectively. For this situation, only the read write head needs to be moved to the left to prepare a initial configuration for TM $adjust$ to start with. \end{enumerate} The diagram of $mainwork$ is given in Figure \ref{mainwork_diag}. The two rectangular nodes labeled with $2 \times x$ and $4 \times x$ are two TMs compiling from recursive functions so that we do not have to design and verify two quite complicated TMs. \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$1$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$}; \node (12) [right = -0.9pt of 11] {\ldots \ldots}; \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {$0$}; \node (14) [draw, text height = 3.9pt, right = -0.9pt of 13, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13); \end{tikzpicture}} \caption{The first situation for TM $mainwork$ to consider} \label{mainwork_case_one_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$}; \node (12) [right = -0.9pt of 11] {\ldots \ldots}; \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {$0$}; \node (14) [draw, text height = 2.7pt, right = -0.9pt of 13, thick, inner sep = 5pt] {$(r+1) \times 2$}; \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13); \end{tikzpicture}} \caption{The output for the first case of TM $mainwork$'s processing} \label{mainwork_case_one_output} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$1$}; \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {$0$}; \node (13) [right = -0.9pt of 12] {\ldots \ldots}; \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {$0$}; \node (15) [draw, text height = 3.9pt, right = -0.9pt of 14, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14); \end{tikzpicture}} \caption{The second situation for TM $mainwork$ to consider} \label{mainwork_case_two_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$}; \node (7) [right = -0.9pt of 6] {\ldots \ldots}; \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$}; \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$}; \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$}; \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$}; \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {$0$}; \node (13) [right = -0.9pt of 12] {\ldots \ldots}; \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {$0$}; \node (15) [draw, text height = 2.7pt, right = -0.9pt of 14, thick, inner sep = 5pt] {$(r+1) \times 4$}; \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14); \end{tikzpicture}} \caption{The output for the second case of TM $mainwork$'s processing} \label{mainwork_case_two_output} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$}; \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (7)+(0, -4\baseheight) -- (7); \end{tikzpicture}} \caption{The third situation for TM $mainwork$ to consider} \label{mainwork_case_three_input} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$}; \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3); \end{tikzpicture}} \caption{The output for the third case of TM $mainwork$'s processing} \label{mainwork_case_three_output} \end{figure} \begin{figure}[h!] \centering \scalebox{0.9}{ \begin{tikzpicture} \node[circle,draw] (1) {$1$}; \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$}; \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$}; \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$}; \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$}; \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$}; \node[circle,draw] (7) at ($(2)+(0, -7\baseheight)$) {$7$}; \node[circle,draw] (8) at ($(7)+(0, -7\baseheight)$) {$8$}; \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$}; \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$}; \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$}; \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$12$}; \node[draw] (13) at ($(6)+(0.3\basewidth, 0)$) {$2 \times x$}; \node[circle,draw] (14) at ($(13)+(0.3\basewidth, 0)$) {$j_1$}; \node[draw] (15) at ($(12)+(0.3\basewidth, 0)$) {$4 \times x$}; \node[draw] (16) at ($(15)+(0.3\basewidth, 0)$) {$j_2$}; \node[draw] (17) at ($(7)+(0.3\basewidth, 0)$) {$0$}; \draw [->, >=latex] (1) edge[loop left] node[above] {$S_0:L$} (1) ; \draw [->, >=latex] (1) -- node[above] {$S_1:L$} (2) ; \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3) ; \draw [->, >=latex] (2) -- node[left] {$S_1:L$} (7) ; \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3) ; \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4) ; \draw [->, >=latex] (4) -- node[above] {$S_1:R$} (5) ; \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5) ; \draw [->, >=latex] (5) -- node[above] {$S_0:S_1$} (6) ; \draw [->, >=latex] (6) edge[loop above] node[above] {$S_1:L$} (6) ; \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (13) ; \draw [->, >=latex] (13) -- (14) ; \draw (14) -- ($(14)+(0, 6\baseheight)$) -- ($(1) + (0, 6\baseheight)$) node [above,midway] {$S_1:L$} ; \draw [->, >=latex] ($(1) + (0, 6\baseheight)$) -- (1) ; \draw [->, >=latex] (7) -- node[above] {$S_0:R$} (17) ; \draw [->, >=latex] (7) -- node[left] {$S_1:R$} (8) ; \draw [->, >=latex] (8) -- node[above] {$S_0:R$} (9) ; \draw [->, >=latex] (9) -- node[above] {$S_0:R$} (10) ; \draw [->, >=latex] (10) -- node[above] {$S_1:R$} (11) ; \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:R$} (10) ; \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:R$} (11) ; \draw [->, >=latex] (11) -- node[above] {$S_0:S_1$} (12) ; \draw [->, >=latex] (12) -- node[above] {$S_0:R$} (15) ; \draw [->, >=latex] (12) edge[loop above] node[above] {$S_1:L$} (12) ; \draw [->, >=latex] (15) -- (16) ; \draw (16) -- ($(16)+(0, -4\baseheight)$) -- ($(1) + (0, -18\baseheight)$) node [below,midway] {$S_1:L$} ; \draw [->, >=latex] ($(1) + (0, -18\baseheight)$) -- (1) ; \end{tikzpicture}} \caption{The diagram of TM $mainwork$} \label{mainwork_diag} \end{figure} The purpose of TM $adjust$ is to encode the last bit of $a_1$. The initial and final configuration of this TM are shown in Figure \ref{adjust_initial} and \ref{adjust_final} respectively. The diagram of TM $adjust$ is shown in Figure \ref{adjust_diag}. \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$}; \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$}; \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3); \end{tikzpicture}} \caption{Initial configuration of TM $adjust$} \label{adjust_initial} \end{figure} \begin{figure}[h!] \centering \scalebox{1.2}{ \begin{tikzpicture} \node (0) {}; \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$}; \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$}; \node (3) [draw, text height = 2.9pt, right = -0.9pt of 2, thick, inner sep = 5pt] {$r+1$}; \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$0$}; \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$}; \node (6) [right = -0.9pt of 5] {\ldots \ldots}; \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1); \end{tikzpicture}} \caption{Final configuration of TM $adjust$} \label{adjust_final} \end{figure} \begin{figure}[h!] \centering \scalebox{0.9}{ \begin{tikzpicture} \node[circle,draw] (1) {$1$}; \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$}; \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$}; \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$}; \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$}; \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$}; \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$}; \node[circle,draw] (8) at ($(4)+(0, -7\baseheight)$) {$8$}; \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$}; \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$}; \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$}; \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$0$}; \draw [->, >=latex] (1) -- node[above] {$S_1:R$} (2) ; \draw [->, >=latex] (1) edge[loop above] node[above] {$S_0:S_1$} (1) ; \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3) ; \draw [->, >=latex] (3) edge[loop above] node[above] {$S_0:R$} (3) ; \draw [->, >=latex] (3) -- node[above] {$S_1:R$} (4) ; \draw [->, >=latex] (4) -- node[above] {$S_1:L$} (5) ; \draw [->, >=latex] (4) -- node[right] {$S_0:L$} (8) ; \draw [->, >=latex] (5) -- node[above] {$S_0:L$} (6) ; \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:S_0$} (5) ; \draw [->, >=latex] (6) -- node[above] {$S_1:R$} (7) ; \draw [->, >=latex] (6) edge[loop above] node[above] {$S_0:L$} (6) ; \draw (7) -- ($(7)+(0, 6\baseheight)$) -- ($(2) + (0, 6\baseheight)$) node [above,midway] {$S_0:S_1$} ; \draw [->, >=latex] ($(2) + (0, 6\baseheight)$) -- (2) ; \draw [->, >=latex] (8) edge[loop left] node[left] {$S_1:S_0$} (8) ; \draw [->, >=latex] (8) -- node[above] {$S_0:L$} (9) ; \draw [->, >=latex] (9) edge[loop above] node[above] {$S_0:L$} (9) ; \draw [->, >=latex] (9) -- node[above] {$S_1:L$} (10) ; \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:L$} (10) ; \draw [->, >=latex] (10) -- node[above] {$S_0:L$} (11) ; \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:L$} (11) ; \draw [->, >=latex] (11) -- node[above] {$S_0:R$} (12) ; \end{tikzpicture}} \caption{Diagram of TM $adjust$} \label{adjust_diag} \end{figure} \ definition rec_twice :: "recf" where "rec_twice = Cn 1 rec_mult [id 1 0, constn 2]" definition rec_fourtimes :: "recf" where "rec_fourtimes = Cn 1 rec_mult [id 1 0, constn 4]" definition abc_twice :: "abc_prog" where "abc_twice = (let (aprog, ary, fp) = rec_ci rec_twice in aprog [+] dummy_abc ((Suc 0)))" definition abc_fourtimes :: "abc_prog" where "abc_fourtimes = (let (aprog, ary, fp) = rec_ci rec_fourtimes in aprog [+] dummy_abc ((Suc 0)))" definition twice_ly :: "nat list" where "twice_ly = layout_of abc_twice" definition fourtimes_ly :: "nat list" where "fourtimes_ly = layout_of abc_fourtimes" definition twice_compile_tm :: "instr list" where "twice_compile_tm= (tm_of abc_twice @ (shift (mopup_n_tm 1) (length (tm_of abc_twice) div 2)))" definition twice_tm :: "instr list" where "twice_tm = adjust0 twice_compile_tm" definition fourtimes_compile_tm :: "instr list" where "fourtimes_compile_tm= (tm_of abc_fourtimes @ (shift (mopup_n_tm 1) (length (tm_of abc_fourtimes) div 2)))" definition fourtimes_tm :: "instr list" where "fourtimes_tm = adjust0 fourtimes_compile_tm" definition twice_tm_len :: "nat" where "twice_tm_len = length twice_tm div 2" definition wcode_main_first_part_tm:: "instr list" where "wcode_main_first_part_tm \ [(L, 1), (L, 2), (L, 7), (R, 3), (R, 4), (WB, 3), (R, 4), (R, 5), (WO, 6), (R, 5), (R, 13), (L, 6), (R, 0), (R, 8), (R, 9), (Nop, 8), (R, 10), (WB, 9), (R, 10), (R, 11), (WO, 12), (R, 11), (R, twice_tm_len + 14), (L, 12)]" definition wcode_main_tm :: "instr list" where "wcode_main_tm = (wcode_main_first_part_tm @ shift twice_tm 12 @ [(L, 1), (L, 1)] @ shift fourtimes_tm (twice_tm_len + 13) @ [(L, 1), (L, 1)])" fun bl_bin :: "cell list \ nat" where "bl_bin [] = 0" | "bl_bin (Bk # xs) = 2 * bl_bin xs" | "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)" declare bl_bin.simps[simp del] type_synonym bin_inv_t = "cell list \ nat \ tape \ bool" fun wcode_before_double :: "bin_inv_t" where "wcode_before_double ires rs (l, r) = (\ ln rn. l = Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\((Suc (Suc rs))) @ Bk\(rn ))" declare wcode_before_double.simps[simp del] fun wcode_after_double :: "bin_inv_t" where "wcode_after_double ires rs (l, r) = (\ ln rn. l = Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(Suc (Suc (Suc 2*rs))) @ Bk\(rn))" declare wcode_after_double.simps[simp del] fun wcode_on_left_moving_1_B :: "bin_inv_t" where "wcode_on_left_moving_1_B ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0 \ mr > 0)" declare wcode_on_left_moving_1_B.simps[simp del] fun wcode_on_left_moving_1_O :: "bin_inv_t" where "wcode_on_left_moving_1_O ires rs (l, r) = (\ ln rn. l = Oc # ires \ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" declare wcode_on_left_moving_1_O.simps[simp del] fun wcode_on_left_moving_1 :: "bin_inv_t" where "wcode_on_left_moving_1 ires rs (l, r) = (wcode_on_left_moving_1_B ires rs (l, r) \ wcode_on_left_moving_1_O ires rs (l, r))" declare wcode_on_left_moving_1.simps[simp del] fun wcode_on_checking_1 :: "bin_inv_t" where "wcode_on_checking_1 ires rs (l, r) = (\ ln rn. l = ires \ r = Oc # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_erase1 :: "bin_inv_t" where "wcode_erase1 ires rs (l, r) = (\ ln rn. l = Oc # ires \ tl r = Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" declare wcode_erase1.simps [simp del] fun wcode_on_right_moving_1 :: "bin_inv_t" where "wcode_on_right_moving_1 ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0)" declare wcode_on_right_moving_1.simps [simp del] (* declare wcode_on_right_moving_1.simps[simp del] *) fun wcode_goon_right_moving_1 :: "bin_inv_t" where "wcode_goon_right_moving_1 ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs)" declare wcode_goon_right_moving_1.simps[simp del] fun wcode_backto_standard_pos_B :: "bin_inv_t" where "wcode_backto_standard_pos_B ires rs (l, r) = (\ ln rn. l = Bk # Bk\(ln) @ Oc # ires \ r = Bk # Oc\((Suc (Suc rs))) @ Bk\(rn ))" declare wcode_backto_standard_pos_B.simps[simp del] fun wcode_backto_standard_pos_O :: "bin_inv_t" where "wcode_backto_standard_pos_O ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0)" declare wcode_backto_standard_pos_O.simps[simp del] fun wcode_backto_standard_pos :: "bin_inv_t" where "wcode_backto_standard_pos ires rs (l, r) = (wcode_backto_standard_pos_B ires rs (l, r) \ wcode_backto_standard_pos_O ires rs (l, r))" declare wcode_backto_standard_pos.simps[simp del] lemma bin_wc_eq: "bl_bin xs = bl2wc xs" proof(induct xs) show " bl_bin [] = bl2wc []" apply(simp add: bl_bin.simps) done next fix a xs assume "bl_bin xs = bl2wc xs" thus " bl_bin (a # xs) = bl2wc (a # xs)" apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps) apply(simp_all add: bl2nat.simps bl2nat_double) done qed lemma tape_of_nl_append_one: "lm \ [] \ = @ Bk # Oc\Suc a" apply(induct lm, auto simp: tape_of_nl_cons split:if_splits) done lemma tape_of_nl_rev: "rev () = ()" apply(induct lm, simp, auto) apply(auto simp: tape_of_nl_cons tape_of_nl_append_one split: if_splits) apply(simp add: exp_ind[THEN sym]) done lemma exp_1[simp]: "a\(Suc 0) = [a]" by(simp) lemma tape_of_nl_cons_app1: "() = (Oc\(Suc a) @ Bk # ())" apply(case_tac xs; simp add: tape_of_list_def tape_of_nat_def) done lemma bl_bin_bk_oc[simp]: "bl_bin (xs @ [Bk, Oc]) = bl_bin xs + 2*2^(length xs)" apply(simp add: bin_wc_eq) using bl2nat_cons_oc[of "xs @ [Bk]"] apply(simp add: bl2nat_cons_bk bl2wc.simps) done lemma tape_of_nat[simp]: "() = Oc\(Suc a)" apply(simp add: tape_of_nat_def) done lemma tape_of_nl_cons_app2: "() = ( @ Bk # Oc\(Suc b))" proof(induct "length xs" arbitrary: xs c, simp add: tape_of_list_def) fix x xs c assume ind: "\xs c. x = length xs \ = @ Bk # Oc\(Suc b)" and h: "Suc x = length (xs::nat list)" show " = @ Bk # Oc\(Suc b)" proof(cases xs, simp add: tape_of_list_def) fix a list assume g: "xs = a # list" hence k: " = @ Bk # Oc\(Suc b)" apply(rule_tac ind) using h apply(simp) done from g and k show " = @ Bk # Oc\(Suc b)" apply(simp add: tape_of_list_def) done qed qed lemma length_2_elems[simp]: "length () = Suc (Suc aa) + length ()" apply(simp add: tape_of_list_def) done lemma bl_bin_addition[simp]: "bl_bin (Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) = bl_bin (Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista)) + 2* 2^(length (Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista)))" using bl_bin_bk_oc[of "Oc\(Suc aa) @ Bk # tape_of_nat_list (a # lista)"] apply(simp) done declare replicate_Suc[simp del] lemma bl_bin_2[simp]: "bl_bin () + (4 * rs + 4) * 2 ^ (length () - Suc 0) = bl_bin (Oc\(Suc aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))" apply(case_tac "list", simp add: add_mult_distrib) apply(simp add: tape_of_nl_cons_app2 add_mult_distrib) apply(simp add: tape_of_list_def) done lemma tape_of_nl_app_Suc: "(()) = () @ [Oc]" proof(induct list) case (Cons a list) then show ?case by(cases list;simp_all add:tape_of_list_def exp_ind) qed (simp add: tape_of_list_def exp_ind) lemma bl_bin_3[simp]: "bl_bin (Oc # Oc\(aa) @ Bk # @ [Oc]) = bl_bin (Oc # Oc\(aa) @ Bk # ) + 2^(length (Oc # Oc\(aa) @ Bk # ))" apply(simp add: bin_wc_eq) apply(simp add: bl2nat_cons_oc bl2wc.simps) using bl2nat_cons_oc[of "Oc # Oc\(aa) @ Bk # "] apply(simp) done lemma bl_bin_4[simp]: "bl_bin (Oc # Oc\(aa) @ Bk # ) + (4 * 2 ^ (aa + length ()) + 4 * (rs * 2 ^ (aa + length ()))) = bl_bin (Oc # Oc\(aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))" apply(simp add: tape_of_nl_app_Suc) done declare tape_of_nat[simp del] fun wcode_double_case_inv :: "nat \ bin_inv_t" where "wcode_double_case_inv st ires rs (l, r) = (if st = Suc 0 then wcode_on_left_moving_1 ires rs (l, r) else if st = Suc (Suc 0) then wcode_on_checking_1 ires rs (l, r) else if st = 3 then wcode_erase1 ires rs (l, r) else if st = 4 then wcode_on_right_moving_1 ires rs (l, r) else if st = 5 then wcode_goon_right_moving_1 ires rs (l, r) else if st = 6 then wcode_backto_standard_pos ires rs (l, r) else if st = 13 then wcode_before_double ires rs (l, r) else False)" declare wcode_double_case_inv.simps[simp del] fun wcode_double_case_state :: "config \ nat" where "wcode_double_case_state (st, l, r) = 13 - st" fun wcode_double_case_step :: "config \ nat" where "wcode_double_case_step (st, l, r) = (if st = Suc 0 then (length l) else if st = Suc (Suc 0) then (length r) else if st = 3 then if hd r = Oc then 1 else 0 else if st = 4 then (length r) else if st = 5 then (length r) else if st = 6 then (length l) else 0)" fun wcode_double_case_measure :: "config \ nat \ nat" where "wcode_double_case_measure (st, l, r) = (wcode_double_case_state (st, l, r), wcode_double_case_step (st, l, r))" definition wcode_double_case_le :: "(config \ config) set" where "wcode_double_case_le \ (inv_image lex_pair wcode_double_case_measure)" lemma wf_lex_pair[intro]: "wf lex_pair" by(auto simp:lex_pair_def) lemma wf_wcode_double_case_le[intro]: "wf wcode_double_case_le" by(auto simp: wcode_double_case_le_def ) lemma fetch_wcode_main_tm[simp]: "fetch wcode_main_tm (Suc 0) Bk = (L, Suc 0)" "fetch wcode_main_tm (Suc 0) Oc = (L, Suc (Suc 0))" "fetch wcode_main_tm (Suc (Suc 0)) Oc = (R, 3)" "fetch wcode_main_tm (Suc (Suc 0)) Bk = (L, 7)" "fetch wcode_main_tm (Suc (Suc (Suc 0))) Bk = (R, 4)" "fetch wcode_main_tm (Suc (Suc (Suc 0))) Oc = (WB, 3)" "fetch wcode_main_tm 4 Bk = (R, 4)" "fetch wcode_main_tm 4 Oc = (R, 5)" "fetch wcode_main_tm 5 Oc = (R, 5)" "fetch wcode_main_tm 5 Bk = (WO, 6)" "fetch wcode_main_tm 6 Bk = (R, 13)" "fetch wcode_main_tm 6 Oc = (L, 6)" "fetch wcode_main_tm 7 Oc = (R, 8)" "fetch wcode_main_tm 7 Bk = (R, 0)" "fetch wcode_main_tm 8 Bk = (R, 9)" "fetch wcode_main_tm 9 Bk = (R, 10)" "fetch wcode_main_tm 9 Oc = (WB, 9)" "fetch wcode_main_tm 10 Bk = (R, 10)" "fetch wcode_main_tm 10 Oc = (R, 11)" "fetch wcode_main_tm 11 Bk = (WO, 12)" "fetch wcode_main_tm 11 Oc = (R, 11)" "fetch wcode_main_tm 12 Oc = (L, 12)" "fetch wcode_main_tm 12 Bk = (R, twice_tm_len + 14)" by(auto simp: wcode_main_tm_def wcode_main_first_part_tm_def fetch.simps numeral_eqs_upto_12) declare wcode_on_checking_1.simps[simp del] lemmas wcode_double_case_inv_simps = wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps wcode_on_left_moving_1_B.simps wcode_on_checking_1.simps wcode_erase1.simps wcode_on_right_moving_1.simps wcode_goon_right_moving_1.simps wcode_backto_standard_pos.simps lemma wcode_on_left_moving_1[simp]: "wcode_on_left_moving_1 ires rs (b, []) = False" "wcode_on_left_moving_1 ires rs (b, r) \ b \ []" by(auto simp: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps wcode_on_left_moving_1_O.simps) lemma wcode_on_left_moving_1E[elim]: "\wcode_on_left_moving_1 ires rs (b, Bk # list); tl b = aa \ hd b # Bk # list = ba\ \ wcode_on_left_moving_1 ires rs (aa, ba)" apply(simp only: wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps wcode_on_left_moving_1_B.simps) apply(erule_tac disjE) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(case_tac ml, simp) apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI) - apply (smt One_nat_def Suc_diff_Suc append_Cons empty_replicate list.sel(3) neq0_conv replicate_Suc + apply (smt (verit) One_nat_def Suc_diff_Suc append_Cons empty_replicate list.sel(3) neq0_conv replicate_Suc replicate_app_Cons_same tl_append2 tl_replicate) apply(rule_tac disjI1) apply (metis add_Suc_shift less_SucI list.exhaust_sel list.inject list.simps(3) replicate_Suc_iff_anywhere) by simp declare replicate_Suc[simp] lemma wcode_on_moving_1_Elim[elim]: "\wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \ hd b # Oc # list = ba\ \ wcode_on_checking_1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac disjE) apply (metis cell.distinct(1) empty_replicate hd_append2 hd_replicate list.sel(1) not_gr_zero) apply force. lemma wcode_on_checking_1_Elim[elim]: "\wcode_on_checking_1 ires rs (b, Oc # ba);Oc # b = aa \ list = ba\ \ wcode_erase1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac exE)+ by auto lemma wcode_on_checking_1_simp[simp]: "wcode_on_checking_1 ires rs (b, []) = False" "wcode_on_checking_1 ires rs (b, Bk # list) = False" by(auto simp: wcode_double_case_inv_simps) lemma wcode_erase1_nonempty_snd[simp]: "wcode_erase1 ires rs (b, []) = False" apply(simp add: wcode_double_case_inv_simps) done lemma wcode_on_right_moving_1_nonempty_snd[simp]: "wcode_on_right_moving_1 ires rs (b, []) = False" apply(simp add: wcode_double_case_inv_simps) done lemma wcode_on_right_moving_1_BkE[elim]: "\wcode_on_right_moving_1 ires rs (b, Bk # ba); Bk # b = aa \ list = b\ \ wcode_on_right_moving_1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, rule_tac x = rn in exI) apply(simp) apply(case_tac mr, simp, simp) done lemma wcode_on_right_moving_1_OcE[elim]: "\wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \ list = ba\ \ wcode_goon_right_moving_1 ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI, rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI) apply(case_tac mr, simp_all) apply(case_tac ml, simp, case_tac nat, simp, simp) done lemma wcode_erase1_BkE[elim]: assumes "wcode_erase1 ires rs (b, Bk # ba)" "Bk # b = aa \ list = ba" "c = Bk # ba" shows "wcode_on_right_moving_1 ires rs (aa, ba)" proof - from assms obtain rn ln where "b = Oc # ires" "tl (Bk # ba) = Bk \ ln @ Bk # Bk # Oc \ Suc rs @ Bk \ rn" unfolding wcode_double_case_inv_simps by auto thus ?thesis using assms(2-) unfolding wcode_double_case_inv_simps apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI, rule_tac x = rn in exI, simp add: exp_ind del: replicate_Suc) done qed lemma wcode_erase1_OcE[elim]: "\wcode_erase1 ires rs (aa, Oc # list); b = aa \ Bk # list = ba\ \ wcode_erase1 ires rs (aa, ba)" unfolding wcode_double_case_inv_simps by auto auto lemma wcode_goon_right_moving_1_emptyE[elim]: assumes "wcode_goon_right_moving_1 ires rs (aa, [])" "b = aa \ [Oc] = ba" shows "wcode_backto_standard_pos ires rs (aa, ba)" proof - from assms obtain ml ln rn mr where "aa = Oc \ ml @ Bk # Bk # Bk \ ln @ Oc # ires" "[] = Oc \ mr @ Bk \ rn" "ml + mr = Suc rs" by(auto simp:wcode_double_case_inv_simps) thus ?thesis using assms(2) apply(simp only: wcode_double_case_inv_simps) apply(rule_tac disjI2) apply(simp only:wcode_backto_standard_pos_O.simps) apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp) done qed lemma wcode_goon_right_moving_1_BkE[elim]: assumes "wcode_goon_right_moving_1 ires rs (aa, Bk # list)" "b = aa \ Oc # list = ba" shows "wcode_backto_standard_pos ires rs (aa, ba)" proof - from assms obtain ln rn where "aa = Oc \ Suc rs @ Bk \ Suc (Suc ln) @ Oc # ires" "Bk # list = Bk \ rn" "b = Oc \ Suc rs @ Bk \ Suc (Suc ln) @ Oc # ires" "ba = Oc # list" by(auto simp:wcode_double_case_inv_simps) thus ?thesis using assms(2) apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps) apply(rule_tac disjI2) apply(rule exI[of _ "Suc rs"], rule exI[of _ "Suc 0"], rule_tac x = ln in exI, rule_tac x = "rn - Suc 0" in exI, simp) apply(cases rn;auto) done qed lemma wcode_goon_right_moving_1_OcE[elim]: assumes "wcode_goon_right_moving_1 ires rs (b, Oc # ba)" "Oc # b = aa \ list = ba" shows "wcode_goon_right_moving_1 ires rs (aa, ba)" proof - from assms obtain ml mr ln rn where "b = Oc \ ml @ Bk # Bk # Bk \ ln @ Oc # ires \ Oc # ba = Oc \ mr @ Bk \ rn \ ml + mr = Suc rs" unfolding wcode_double_case_inv_simps by auto with assms(2) show ?thesis unfolding wcode_double_case_inv_simps apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) apply(simp) apply(case_tac mr, simp, case_tac rn, simp_all) done qed lemma wcode_backto_standard_pos_BkE[elim]: "\wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \ list = ba\ \ wcode_before_double ires rs (aa, ba)" apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps wcode_before_double.simps) apply(erule_tac disjE) apply(erule_tac exE)+ by auto lemma wcode_backto_standard_pos_no_Oc[simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False" apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps) done lemma wcode_backto_standard_pos_nonempty_snd[simp]: "wcode_backto_standard_pos ires rs (b, []) = False" apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps) done lemma wcode_backto_standard_pos_OcE[elim]: "\wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list = ba\ \ wcode_backto_standard_pos ires rs (aa, ba)" apply(simp only: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps wcode_backto_standard_pos_O.simps) apply(erule_tac disjE) apply(simp) apply(erule_tac exE)+ apply(simp) apply (rename_tac ml mr ln rn) apply(case_tac ml) apply(rule_tac disjI1, rule_tac conjI) apply(rule_tac x = ln in exI, force, rule_tac x = rn in exI, force, force). declare nth_of.simps[simp del] lemma wcode_double_case_first_correctness: "let P = (\ (st, l, r). st = 13) in let Q = (\ (st, l, r). wcode_double_case_inv st ires rs (l, r)) in let f = (\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp) in \ n .P (f n) \ Q (f (n::nat))" proof - let ?P = "(\ (st, l, r). st = 13)" let ?Q = "(\ (st, l, r). wcode_double_case_inv st ires rs (l, r))" let ?f = "(\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp)" have "\ n. ?P (?f n) \ ?Q (?f (n::nat))" proof(rule_tac halt_lemma2) show "wf wcode_double_case_le" by auto next show "\ na. \ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_double_case_le" proof(rule_tac allI, case_tac "?f na", simp) fix na a b c show "a \ 13 \ wcode_double_case_inv a ires rs (b, c) \ (case step0 (a, b, c) wcode_main_tm of (st, x) \ wcode_double_case_inv st ires rs x) \ (step0 (a, b, c) wcode_main_tm, a, b, c) \ wcode_double_case_le" apply(rule_tac impI, simp add: wcode_double_case_inv.simps) apply(auto split: if_splits simp: step.simps, case_tac [!] c, simp_all, case_tac [!] "(c::cell list)!0") apply(simp_all add: wcode_double_case_inv.simps wcode_double_case_le_def lex_pair_def) apply(auto split: if_splits) done qed next show "?Q (?f 0)" apply(simp add: steps.simps wcode_double_case_inv.simps wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps) apply(rule_tac disjI1) apply(rule_tac x = "Suc m" in exI, simp) apply(rule_tac x = "Suc 0" in exI, simp) done next show "\ ?P (?f 0)" apply(simp add: steps.simps) done qed thus "let P = \(st, l, r). st = 13; Q = \(st, l, r). wcode_double_case_inv st ires rs (l, r); f = steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm in \n. P (f n) \ Q (f n)" apply(simp) done qed lemma tm_append_shift_append_steps: "\steps0 (st, l, r) tp stp = (st', l', r'); 0 < st'; length tp1 mod 2 = 0 \ \ steps0 (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2) @ tp2) stp = (st' + length tp1 div 2, l', r')" proof - assume h: "steps0 (st, l, r) tp stp = (st', l', r')" "0 < st'" "length tp1 mod 2 = 0 " from h have "steps (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2), 0) stp = (st' + length tp1 div 2, l', r')" by(rule_tac tm_append_second_steps_eq, simp_all) then have "steps (st + length tp1 div 2, l, r) ((tp1 @ shift tp (length tp1 div 2)) @ tp2, 0) stp = (st' + length tp1 div 2, l', r')" using h apply(rule_tac tm_append_first_steps_eq, simp_all) done thus "?thesis" by simp qed lemma twice_lemma: "rec_exec rec_twice [rs] = 2*rs" by(auto simp: rec_twice_def rec_exec.simps) lemma twice_tm_correct: "\stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_twice @ shift (mopup_n_tm (Suc 0)) ((length (tm_of abc_twice) div 2))) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" proof(case_tac "rec_ci rec_twice") fix a b c assume h: "rec_ci rec_twice = (a, b, c)" have "\stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\(n)) (tm_of abc_twice @ shift (mopup_n_tm (length [rs])) (length (tm_of abc_twice) div 2)) stp = (0, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (rec_exec rec_twice [rs])) @ Bk\(l))" thm recursive_compile_to_tm_correct1 proof(rule_tac recursive_compile_to_tm_correct1) show "rec_ci rec_twice = (a, b, c)" by (simp add: h) next show "terminate rec_twice [rs]" apply(rule_tac primerec_terminate, auto) apply(simp add: rec_twice_def, auto simp: constn.simps numeral_2_eq_2) by(auto) next show "tm_of abc_twice = tm_of (a [+] dummy_abc (length [rs]))" using h by(simp add: abc_twice_def) qed thus "?thesis" apply(simp add: tape_of_list_def tape_of_nat_def rec_exec.simps twice_lemma) done qed declare adjust.simps[simp] lemma adjust_fetch0: "\0 < a; a \ length ap div 2; fetch ap a b = (aa, 0)\ \ fetch (adjust0 ap) a b = (aa, Suc (length ap div 2))" apply(case_tac b, auto split: if_splits) apply(case_tac [!] a, auto simp: fetch.simps nth_of.simps) done lemma adjust_fetch_norm: "\st > 0; st \ length tp div 2; fetch ap st b = (aa, ns); ns \ 0\ \ fetch (adjust0 ap) st b = (aa, ns)" apply(case_tac b, auto simp: fetch.simps split: if_splits) apply(case_tac [!] st, auto simp: fetch.simps nth_of.simps) done declare adjust.simps[simp del] lemma adjust_step_eq: assumes exec: "step0 (st,l,r) ap = (st', l', r')" and "composable_tm (ap, 0)" and notfinal: "st' > 0" shows "step0 (st, l, r) (adjust0 ap) = (st', l', r')" using assms proof - have "st > 0" using assms by(case_tac st, simp_all add: step.simps fetch.simps) moreover hence "st \ (length ap) div 2" using assms apply(case_tac "st \ (length ap) div 2", simp) apply(case_tac st, auto simp: step.simps fetch.simps) apply(case_tac "read r", simp_all add: fetch.simps nth_of.simps adjust.simps composable_tm.simps split: if_splits) apply(auto simp: mod_ex2) done ultimately have "fetch (adjust0 ap) st (read r) = fetch ap st (read r)" using assms apply(case_tac "fetch ap st (read r)") apply(drule_tac adjust_fetch_norm, simp_all) apply(simp add: step.simps) done thus "?thesis" using exec by(simp add: step.simps) qed lemma adjust_steps_eq: assumes exec: "steps0 (st,l,r) ap stp = (st', l', r')" and "composable_tm (ap, 0)" and notfinal: "st' > 0" shows "steps0 (st, l, r) (adjust0 ap) stp = (st', l', r')" using exec notfinal proof(induct stp arbitrary: st' l' r') case 0 thus "?case" by(simp add: steps.simps) next case (Suc stp st' l' r') have ind: "\st' l' r'. \steps0 (st, l, r) ap stp = (st', l', r'); 0 < st'\ \ steps0 (st, l, r) (adjust0 ap) stp = (st', l', r')" by fact have h: "steps0 (st, l, r) ap (Suc stp) = (st', l', r')" by fact have g: "0 < st'" by fact obtain st'' l'' r'' where a: "steps0 (st, l, r) ap stp = (st'', l'', r'')" by (metis prod_cases3) hence c:"0 < st''" using h g apply(simp) apply(case_tac st'', auto) done hence b: "steps0 (st, l, r) (adjust0 ap) stp = (st'', l'', r'')" using a by(rule_tac ind, simp_all) thus "?case" using assms a b h g apply(simp ) apply(rule_tac adjust_step_eq, simp_all) done qed lemma adjust_halt_eq: assumes exec: "steps0 (1, l, r) ap stp = (0, l', r')" and composable_tm: "composable_tm (ap, 0)" shows "\ stp. steps0 (Suc 0, l, r) (adjust0 ap) stp = (Suc (length ap div 2), l', r')" proof - have "\ stp. \ is_final (steps0 (1, l, r) ap stp) \ (steps0 (1, l, r) ap (Suc stp) = (0, l', r'))" using exec by(erule_tac before_final) then obtain stpa where a: "\ is_final (steps0 (1, l, r) ap stpa) \ (steps0 (1, l, r) ap (Suc stpa) = (0, l', r'))" .. obtain sa la ra where b:"steps0 (1, l, r) ap stpa = (sa, la, ra)" by (metis prod_cases3) hence c: "steps0 (Suc 0, l, r) (adjust0 ap) stpa = (sa, la, ra)" using assms a apply(rule_tac adjust_steps_eq, simp_all) done have d: "sa \ length ap div 2" using steps_in_range[of "(l, r)" ap stpa] a composable_tm b by(simp) obtain ac ns where e: "fetch ap sa (read ra) = (ac, ns)" by (metis prod.exhaust) hence f: "ns = 0" using b a apply(simp add: step.simps) done have k: "fetch (adjust0 ap) sa (read ra) = (ac, Suc (length ap div 2))" using a b c d e f apply(rule_tac adjust_fetch0, simp_all) done from a b e f k and c show "?thesis" apply(rule_tac x = "Suc stpa" in exI) apply(simp , auto) apply(simp add: step.simps) done qed lemma composable_tm_twice_compile_tm [simp]: "composable_tm (twice_compile_tm, 0)" apply(simp only: twice_compile_tm_def) apply(rule_tac composable_tm_from_abacus, simp) done lemma twice_tm_change_term_state: "\ stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) twice_tm stp = (Suc twice_tm_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" proof - have "\stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_twice @ shift (mopup_n_tm (Suc 0)) ((length (tm_of abc_twice) div 2))) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" by(rule_tac twice_tm_correct) then obtain stp ln rn where " steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_twice @ shift (mopup_n_tm (Suc 0)) ((length (tm_of abc_twice) div 2))) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" by blast hence "\ stp. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 twice_compile_tm) stp = (Suc (length twice_compile_tm div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" apply(rule_tac stp = stp in adjust_halt_eq) apply(simp add: twice_compile_tm_def, auto) done then obtain stpb where "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 twice_compile_tm) stpb = (Suc (length twice_compile_tm div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" .. thus "?thesis" apply(simp add: twice_tm_def twice_tm_len_def) by metis qed lemma length_wcode_main_first_part_tm_even[intro]: "length wcode_main_first_part_tm mod 2 = 0" apply(auto simp: wcode_main_first_part_tm_def) done lemma twice_tm_append_pre: "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) twice_tm stp = (Suc twice_tm_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn)) \ steps0 (Suc 0 + length wcode_main_first_part_tm div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ ([(L, 1), (L, 1)] @ shift fourtimes_tm (twice_tm_len + 13) @ [(L, 1), (L, 1)])) stp = (Suc (twice_tm_len) + length wcode_main_first_part_tm div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" by(rule_tac tm_append_shift_append_steps, auto) lemma twice_tm_append: "\ stp ln rn. steps0 (Suc 0 + length wcode_main_first_part_tm div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ ([(L, 1), (L, 1)] @ shift fourtimes_tm (twice_tm_len + 13) @ [(L, 1), (L, 1)])) stp = (Suc (twice_tm_len) + length wcode_main_first_part_tm div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(rn))" using twice_tm_change_term_state[of ires rs n] apply(erule_tac exE) apply(erule_tac exE) apply(erule_tac exE) apply(drule_tac twice_tm_append_pre) apply(rename_tac stp ln rn) apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) apply(simp) done lemma mopup_mod2: "length (mopup_n_tm k) mod 2 = 0" by(auto simp: mopup_n_tm.simps) lemma fetch_wcode_main_tm_Oc[simp]: "fetch wcode_main_tm (Suc (twice_tm_len + length wcode_main_first_part_tm div 2)) Oc = (L, Suc 0)" apply(subgoal_tac "length (twice_tm) mod 2 = 0") apply(simp add: wcode_main_tm_def nth_append fetch.simps wcode_main_first_part_tm_def nth_of.simps twice_tm_len_def, auto) apply(simp add: twice_tm_def twice_compile_tm_def) using mopup_mod2[of 1] apply(simp) done lemma wcode_jump1: "\ stp ln rn. steps0 (Suc (twice_tm_len) + length wcode_main_first_part_tm div 2, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (2 * rs)) @ Bk\(n)) wcode_main_tm stp = (Suc 0, Bk\(ln) @ Bk # ires, Bk # Oc\(Suc (2 * rs)) @ Bk\(rn))" apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI) apply(simp add: steps.simps step.simps exp_ind) apply(case_tac m, simp_all) apply(simp add: exp_ind[THEN sym]) done lemma wcode_main_first_part_len[simp]: "length wcode_main_first_part_tm = 24" apply(simp add: wcode_main_first_part_tm_def) done lemma wcode_double_case: shows "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rn))" (is "\stp ln rn. ?tm stp ln rn") proof - from wcode_double_case_first_correctness[of ires rs m n] obtain na ln rn where "steps0 (Suc 0, Bk # Bk \ m @ Oc # Oc # ires, Bk # Oc # Oc \ rs @ Bk \ n) wcode_main_tm na = (13, Bk # Bk # Bk \ ln @ Oc # ires, Oc # Oc # Oc \ rs @ Bk \ rn)" by(auto simp: wcode_double_case_inv.simps wcode_before_double.simps) hence "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (13, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rn))" by(case_tac "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm na", auto) from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stpa = (13, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rna))" by blast from twice_tm_append[of "Bk\(lna) @ Oc # ires" "Suc rs" rna] obtain stp ln rn where "steps0 (Suc 0 + length wcode_main_first_part_tm div 2, Bk # Bk # Bk \ lna @ Oc # ires, Oc \ Suc (Suc rs) @ Bk \ rna) (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)] @ shift fourtimes_tm (twice_tm_len + 13) @ [(L, 1), (L, 1)]) stp = (Suc twice_tm_len + length wcode_main_first_part_tm div 2, Bk \ ln @ Bk # Bk # Bk \ lna @ Oc # ires, Oc \ Suc (2 * Suc rs) @ Bk \ rn)" by blast hence "\ stp ln rn. steps0 (13, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rna)) wcode_main_tm stp = (13 + twice_tm_len, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rn))" using twice_tm_append[of "Bk\(lna) @ Oc # ires" "Suc rs" rna] apply(simp) apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, rule_tac x = rn in exI) apply(simp add: wcode_main_tm_def) apply(simp add: replicate_Suc[THEN sym] replicate_add [THEN sym] del: replicate_Suc) done from this obtain stpb lnb rnb where stp2: "steps0 (13, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (Suc rs)) @ Bk\(rna)) wcode_main_tm stpb = (13 + twice_tm_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb))" by blast from wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb] obtain stp ln rn where "steps0 (Suc twice_tm_len + length wcode_main_first_part_tm div 2, Bk \ lnb @ Bk # Bk # Oc # ires, Oc \ Suc (2 * Suc rs) @ Bk \ rnb) wcode_main_tm stp = (Suc 0, Bk \ ln @ Bk # Oc # ires, Bk # Oc \ Suc (2 * Suc rs) @ Bk \ rn)" by metis hence "steps0 (13 + twice_tm_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rn))" apply(auto simp add: wcode_main_tm_def) apply(subgoal_tac "Bk\(lnb) @ Bk # Bk # Oc # ires = Bk # Bk # Bk\(lnb) @ Oc # ires", simp) apply(simp add: replicate_Suc[THEN sym] exp_ind[THEN sym] del: replicate_Suc) apply(simp) apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc) done hence "\stp ln rn. steps0 (13 + twice_tm_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rn))" by blast from this obtain stpc lnc rnc where stp3: "steps0 (13 + twice_tm_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnb)) wcode_main_tm stpc = (Suc 0, Bk # Bk\(lnc) @ Oc # ires, Bk # Oc\(Suc (Suc (Suc (2 *rs)))) @ Bk\(rnc))" by blast from stp1 stp2 stp3 have "?tm (stpa + stpb + stpc) lnc rnc" by simp thus "?thesis" by blast qed (* Begin: fourtime_case*) fun wcode_on_left_moving_2_B :: "bin_inv_t" where "wcode_on_left_moving_2_B ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # Bk # Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0 \ mr > 0)" fun wcode_on_left_moving_2_O :: "bin_inv_t" where "wcode_on_left_moving_2_O ires rs (l, r) = (\ ln rn. l = Bk # Oc # ires \ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_on_left_moving_2 :: "bin_inv_t" where "wcode_on_left_moving_2 ires rs (l, r) = (wcode_on_left_moving_2_B ires rs (l, r) \ wcode_on_left_moving_2_O ires rs (l, r))" fun wcode_on_checking_2 :: "bin_inv_t" where "wcode_on_checking_2 ires rs (l, r) = (\ ln rn. l = Oc#ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_goon_checking :: "bin_inv_t" where "wcode_goon_checking ires rs (l, r) = (\ ln rn. l = ires \ r = Oc # Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_right_move :: "bin_inv_t" where "wcode_right_move ires rs (l, r) = (\ ln rn. l = Oc # ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_erase2 :: "bin_inv_t" where "wcode_erase2 ires rs (l, r) = (\ ln rn. l = Bk # Oc # ires \ tl r = Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_on_right_moving_2 :: "bin_inv_t" where "wcode_on_right_moving_2 ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0)" fun wcode_goon_right_moving_2 :: "bin_inv_t" where "wcode_goon_right_moving_2 ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs)" fun wcode_backto_standard_pos_2_B :: "bin_inv_t" where "wcode_backto_standard_pos_2_B ires rs (l, r) = (\ ln rn. l = Bk # Bk\(ln) @ Oc # ires \ r = Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" fun wcode_backto_standard_pos_2_O :: "bin_inv_t" where "wcode_backto_standard_pos_2_O ires rs (l, r) = (\ ml mr ln rn. l = Oc\(ml )@ Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = (Suc (Suc rs)) \ mr > 0)" fun wcode_backto_standard_pos_2 :: "bin_inv_t" where "wcode_backto_standard_pos_2 ires rs (l, r) = (wcode_backto_standard_pos_2_O ires rs (l, r) \ wcode_backto_standard_pos_2_B ires rs (l, r))" fun wcode_before_fourtimes :: "bin_inv_t" where "wcode_before_fourtimes ires rs (l, r) = (\ ln rn. l = Bk # Bk # Bk\(ln) @ Oc # ires \ r = Oc\(Suc (Suc rs)) @ Bk\(rn))" declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del] wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del] wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del] wcode_erase2.simps[simp del] wcode_on_right_moving_2.simps[simp del] wcode_goon_right_moving_2.simps[simp del] wcode_backto_standard_pos_2_B.simps[simp del] wcode_backto_standard_pos_2_O.simps[simp del] wcode_backto_standard_pos_2.simps[simp del] lemmas wcode_fourtimes_invs = wcode_on_left_moving_2_B.simps wcode_on_left_moving_2.simps wcode_on_left_moving_2_O.simps wcode_on_checking_2.simps wcode_goon_checking.simps wcode_right_move.simps wcode_erase2.simps wcode_on_right_moving_2.simps wcode_goon_right_moving_2.simps wcode_backto_standard_pos_2_B.simps wcode_backto_standard_pos_2_O.simps wcode_backto_standard_pos_2.simps fun wcode_fourtimes_case_inv :: "nat \ bin_inv_t" where "wcode_fourtimes_case_inv st ires rs (l, r) = (if st = Suc 0 then wcode_on_left_moving_2 ires rs (l, r) else if st = Suc (Suc 0) then wcode_on_checking_2 ires rs (l, r) else if st = 7 then wcode_goon_checking ires rs (l, r) else if st = 8 then wcode_right_move ires rs (l, r) else if st = 9 then wcode_erase2 ires rs (l, r) else if st = 10 then wcode_on_right_moving_2 ires rs (l, r) else if st = 11 then wcode_goon_right_moving_2 ires rs (l, r) else if st = 12 then wcode_backto_standard_pos_2 ires rs (l, r) else if st = twice_tm_len + 14 then wcode_before_fourtimes ires rs (l, r) else False)" declare wcode_fourtimes_case_inv.simps[simp del] fun wcode_fourtimes_case_state :: "config \ nat" where "wcode_fourtimes_case_state (st, l, r) = 13 - st" fun wcode_fourtimes_case_step :: "config \ nat" where "wcode_fourtimes_case_step (st, l, r) = (if st = Suc 0 then length l else if st = 9 then (if hd r = Oc then 1 else 0) else if st = 10 then length r else if st = 11 then length r else if st = 12 then length l else 0)" fun wcode_fourtimes_case_measure :: "config \ nat \ nat" where "wcode_fourtimes_case_measure (st, l, r) = (wcode_fourtimes_case_state (st, l, r), wcode_fourtimes_case_step (st, l, r))" definition wcode_fourtimes_case_le :: "(config \ config) set" where "wcode_fourtimes_case_le \ (inv_image lex_pair wcode_fourtimes_case_measure)" lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le" by(auto simp: wcode_fourtimes_case_le_def) lemma nonempty_snd [simp]: "wcode_on_left_moving_2 ires rs (b, []) = False" "wcode_on_checking_2 ires rs (b, []) = False" "wcode_goon_checking ires rs (b, []) = False" "wcode_right_move ires rs (b, []) = False" "wcode_erase2 ires rs (b, []) = False" "wcode_on_right_moving_2 ires rs (b, []) = False" "wcode_backto_standard_pos_2 ires rs (b, []) = False" "wcode_on_checking_2 ires rs (b, Oc # list) = False" by(auto simp: wcode_fourtimes_invs) lemma gr1_conv_Suc:"Suc 0 < mr \ (\ nat. mr = Suc (Suc nat))" by presburger lemma wcode_on_left_moving_2[simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \ wcode_on_left_moving_2 ires rs (tl b, hd b # Bk # list)" apply(simp only: wcode_fourtimes_invs) apply(erule_tac disjE) apply(erule_tac exE)+ apply(simp add: gr1_conv_Suc exp_ind replicate_app_Cons_same hd_repeat_cases') apply (auto simp add: gr0_conv_Suc[symmetric] replicate_app_Cons_same hd_repeat_cases') by force+ lemma wcode_goon_checking_via_2 [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) \ wcode_goon_checking ires rs (tl b, hd b # Bk # list)" unfolding wcode_fourtimes_invs by auto lemma wcode_erase2_via_move [simp]: "wcode_right_move ires rs (b, Bk # list) \ wcode_erase2 ires rs (Bk # b, list)" by (auto simp:wcode_fourtimes_invs ) auto lemma wcode_on_right_moving_2_via_erase2[simp]: "wcode_erase2 ires rs (b, Bk # list) \ wcode_on_right_moving_2 ires rs (Bk # b, list)" apply(auto simp:wcode_fourtimes_invs ) apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind) by (metis replicate_Suc_iff_anywhere replicate_app_Cons_same) lemma wcode_on_right_moving_2_move_Bk[simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) \ wcode_on_right_moving_2 ires rs (Bk # b, list)" apply(auto simp: wcode_fourtimes_invs) apply(rename_tac ml mr rn) apply(rule_tac x = "Suc ml" in exI, simp) apply(rule_tac x = "mr - 1" in exI, case_tac mr,auto) done lemma wcode_backto_standard_pos_2_via_right[simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \ wcode_backto_standard_pos_2 ires rs (b, Oc # list)" apply(simp add: wcode_fourtimes_invs, auto) by (metis add.right_neutral add_Suc_shift append_Cons list.sel(3) replicate.simps(1) replicate_Suc replicate_Suc_iff_anywhere self_append_conv2 tl_replicate zero_less_Suc) lemma wcode_on_checking_2_via_left[simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \ wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)" by(auto simp: wcode_fourtimes_invs) lemma wcode_backto_standard_pos_2_empty_via_right[simp]: "wcode_goon_right_moving_2 ires rs (b, []) \ wcode_backto_standard_pos_2 ires rs (b, [Oc])" by (auto simp add: wcode_fourtimes_invs) force lemma wcode_goon_checking_cases[simp]: "wcode_goon_checking ires rs (b, Oc # list) \ (b = [] \ wcode_right_move ires rs ([Oc], list)) \ (b \ [] \ wcode_right_move ires rs (Oc # b, list))" apply(simp only: wcode_fourtimes_invs) apply(erule_tac exE)+ apply(auto) done lemma wcode_right_move_no_Oc[simp]: "wcode_right_move ires rs (b, Oc # list) = False" apply(auto simp: wcode_fourtimes_invs) done lemma wcode_erase2_Bk_via_Oc[simp]: "wcode_erase2 ires rs (b, Oc # list) \ wcode_erase2 ires rs (b, Bk # list)" apply(auto simp: wcode_fourtimes_invs) done lemma wcode_goon_right_moving_2_Oc_move[simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list) \ wcode_goon_right_moving_2 ires rs (Oc # b, list)" apply(auto simp: wcode_fourtimes_invs) apply(rule_tac x = "Suc 0" in exI, auto) apply(rule_tac x = "ml - 2" in exI) apply(case_tac ml, simp, case_tac "ml - 1", simp_all) done lemma wcode_backto_standard_pos_2_exists[simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \ (\ln. b = Bk # Bk\(ln) @ Oc # ires) \ (\rn. list = Oc\(Suc (Suc rs)) @ Bk\(rn))" by(simp add: wcode_fourtimes_invs) lemma wcode_goon_right_moving_2_move_Oc[simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \ wcode_goon_right_moving_2 ires rs (Oc # b, list)" apply(simp only:wcode_fourtimes_invs, auto) apply(rename_tac ml ln mr rn) apply(case_tac mr;force) done lemma wcode_backto_standard_pos_2_Oc_mv_hd[simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \ wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)" apply(simp only: wcode_fourtimes_invs) apply(erule_tac disjE) apply(erule_tac exE)+ apply(rename_tac ml mr ln rn) by (case_tac ml, force,force,force) lemma nonempty_fst[simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \ b \ []" "wcode_on_checking_2 ires rs (b, Bk # list) \ b \ []" "wcode_goon_checking ires rs (b, Bk # list) = False" "wcode_right_move ires rs (b, Bk # list) \ b\ []" "wcode_erase2 ires rs (b, Bk # list) \ b \ []" "wcode_on_right_moving_2 ires rs (b, Bk # list) \ b \ []" "wcode_goon_right_moving_2 ires rs (b, Bk # list) \ b \ []" "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \ b \ []" "wcode_on_left_moving_2 ires rs (b, Oc # list) \ b \ []" "wcode_goon_right_moving_2 ires rs (b, []) \ b \ []" "wcode_erase2 ires rs (b, Oc # list) \ b \ []" "wcode_on_right_moving_2 ires rs (b, Oc # list) \ b \ []" "wcode_goon_right_moving_2 ires rs (b, Oc # list) \ b \ []" "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \ b \ []" by(auto simp: wcode_fourtimes_invs) lemma wcode_fourtimes_case_first_correctness: shows "let P = (\ (st, l, r). st = twice_tm_len + 14) in let Q = (\ (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in let f = (\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp) in \ n .P (f n) \ Q (f (n::nat))" proof - let ?P = "(\ (st, l, r). st = twice_tm_len + 14)" let ?Q = "(\ (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))" let ?f = "(\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp)" have "\ n . ?P (?f n) \ ?Q (?f (n::nat))" proof(rule_tac halt_lemma2) show "wf wcode_fourtimes_case_le" by auto next have "\ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_fourtimes_case_le" for na apply(cases "?f na", rule_tac impI) apply(simp add: step.simps) apply(case_tac "snd (snd (?f na))", simp, case_tac [2] "hd (snd (snd (?f na)))", simp_all) apply(simp_all add: wcode_fourtimes_case_inv.simps wcode_fourtimes_case_le_def lex_pair_def split: if_splits) by(auto simp: wcode_backto_standard_pos_2.simps wcode_backto_standard_pos_2_O.simps wcode_backto_standard_pos_2_B.simps gr0_conv_Suc) thus "\ na. \ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_fourtimes_case_le" by auto next show "?Q (?f 0)" apply(simp add: steps.simps wcode_fourtimes_case_inv.simps) apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps wcode_on_left_moving_2_O.simps) apply(rule_tac x = "Suc m" in exI, simp ) apply(rule_tac x ="Suc 0" in exI, auto) done next show "\ ?P (?f 0)" apply(simp add: steps.simps) done qed thus "?thesis" apply(erule_tac exE, simp) done qed definition fourtimes_tm_len :: "nat" where "fourtimes_tm_len = (length fourtimes_tm div 2)" lemma primerec_rec_fourtimes_1[intro]: "primerec rec_fourtimes (Suc 0)" apply(auto simp: rec_fourtimes_def numeral_4_eq_4 constn.simps) by auto lemma fourtimes_lemma: "rec_exec rec_fourtimes [rs] = 4 * rs" by(simp add: rec_exec.simps rec_fourtimes_def) lemma fourtimes_tm_correct: "\stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup_n_tm 1) (length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" proof(case_tac "rec_ci rec_fourtimes") fix a b c assume h: "rec_ci rec_fourtimes = (a, b, c)" have "\stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup_n_tm (length [rs])) (length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\(m) @ Bk # Bk # ires, Oc\(Suc (rec_exec rec_fourtimes [rs])) @ Bk\(l))" thm recursive_compile_to_tm_correct1 proof(rule_tac recursive_compile_to_tm_correct1) show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h) next show "terminate rec_fourtimes [rs]" apply(rule_tac primerec_terminate) by auto next show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (length [rs]))" using h by(simp add: abc_fourtimes_def) qed thus "?thesis" apply(simp add: tape_of_list_def tape_of_nat_def fourtimes_lemma) done qed lemma composable_fourtimes_tm[intro]: "composable_tm (fourtimes_compile_tm, 0)" apply(simp only: fourtimes_compile_tm_def) apply(rule_tac composable_tm_from_abacus, simp) done lemma fourtimes_tm_change_term_state: "\ stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) fourtimes_tm stp = (Suc fourtimes_tm_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" proof - have "\stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup_n_tm 1) ((length (tm_of abc_fourtimes) div 2))) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" by(rule_tac fourtimes_tm_correct) then obtain stp ln rn where "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (tm_of abc_fourtimes @ shift (mopup_n_tm 1) ((length (tm_of abc_fourtimes) div 2))) stp = (0, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" by blast hence "\ stp. steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 fourtimes_compile_tm) stp = (Suc (length fourtimes_compile_tm div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" apply(rule_tac stp = stp in adjust_halt_eq) apply(simp add: fourtimes_compile_tm_def, auto) done then obtain stpb where "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) (adjust0 fourtimes_compile_tm) stpb = (Suc (length fourtimes_compile_tm div 2), Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" .. thus "?thesis" apply(simp add: fourtimes_tm_def fourtimes_tm_len_def) by metis qed lemma length_twice_tm_even[intro]: "is_even (length twice_tm)" by(auto simp: twice_tm_def twice_compile_tm_def intro!:mopup_mod2) lemma fourtimes_tm_append_pre: "steps0 (Suc 0, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) fourtimes_tm stp = (Suc fourtimes_tm_len, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn)) \ steps0 (Suc 0 + length (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)]) div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) ((wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)]) @ shift fourtimes_tm (length (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp = ((Suc fourtimes_tm_len) + length (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)]) div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" using length_twice_tm_even by(intro tm_append_shift_append_steps, auto) lemma split_26_even[simp]: "(26 + l::nat) div 2 = l div 2 + 13" by(simp) lemma twice_tm_len_plus_14[simp]: "twice_tm_len + 14 = 14 + length (shift twice_tm 12) div 2" apply(simp add: twice_tm_def twice_tm_len_def) done lemma fourtimes_tm_append: "\ stp ln rn. steps0 (Suc 0 + length (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)]) div 2, Bk # Bk # ires, Oc\(Suc rs) @ Bk\(n)) ((wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)]) @ shift fourtimes_tm (twice_tm_len + 13) @ [(L, 1), (L, 1)]) stp = (Suc fourtimes_tm_len + length (wcode_main_first_part_tm @ shift twice_tm (length wcode_main_first_part_tm div 2) @ [(L, 1), (L, 1)]) div 2, Bk\(ln) @ Bk # Bk # ires, Oc\(Suc (4 * rs)) @ Bk\(rn))" using fourtimes_tm_change_term_state[of ires rs n] apply(erule_tac exE) apply(erule_tac exE) apply(erule_tac exE) apply(drule_tac fourtimes_tm_append_pre) apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) apply(simp add: twice_tm_len_def) done lemma even_fourtimes_len: "length fourtimes_tm mod 2 = 0" apply(auto simp: fourtimes_tm_def fourtimes_compile_tm_def) by (metis mopup_mod2) lemma twice_tm_even[simp]: "2 * (length twice_tm div 2) = length twice_tm" using length_twice_tm_even by arith lemma fourtimes_tm_even[simp]: "2 * (length fourtimes_tm div 2) = length fourtimes_tm" using even_fourtimes_len by arith lemma fetch_wcode_tm_14_Oc: "fetch wcode_main_tm (14 + length twice_tm div 2 + fourtimes_tm_len) Oc = (L, Suc 0)" apply(subgoal_tac "14 = Suc 13") apply(simp only: fetch.simps add_Suc nth_of.simps wcode_main_tm_def) apply(simp add:length_twice_tm_even fourtimes_tm_len_def nth_append) by arith lemma fetch_wcode_tm_14_Bk: "fetch wcode_main_tm (14 + length twice_tm div 2 + fourtimes_tm_len) Bk = (L, Suc 0)" apply(subgoal_tac "14 = Suc 13") apply(simp only: fetch.simps add_Suc nth_of.simps wcode_main_tm_def) apply(simp add:length_twice_tm_even fourtimes_tm_len_def nth_append) by arith lemma fetch_wcode_tm_14 [simp]: "fetch wcode_main_tm (14 + length twice_tm div 2 + fourtimes_tm_len) b = (L, Suc 0)" apply(case_tac b, simp_all add:fetch_wcode_tm_14_Bk fetch_wcode_tm_14_Oc) done lemma wcode_jump2: "\ stp ln rn. steps0 (twice_tm_len + 14 + fourtimes_tm_len , Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4 * rs + 4)) @ Bk\(rnb)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (4 * rs + 4)) @ Bk\(rn))" apply(rule_tac x = "Suc 0" in exI) apply(simp add: steps.simps) apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI) apply(simp add: step.simps) done lemma wcode_fourtimes_case: shows "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rn))" proof - have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (twice_tm_len + 14, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rn))" using wcode_fourtimes_case_first_correctness[of ires rs m n] by (auto simp add: wcode_fourtimes_case_inv.simps) auto from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Oc # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stpa = (twice_tm_len + 14, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rna))" by blast have "\stp ln rn. steps0 (twice_tm_len + 14, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rna)) wcode_main_tm stp = (twice_tm_len + 14 + fourtimes_tm_len, Bk # Bk # Bk\(ln) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rn))" using fourtimes_tm_append[of " Bk\(lna) @ Oc # ires" "rs + 1" rna] apply(erule_tac exE) apply(erule_tac exE) apply(erule_tac exE) apply(simp add: wcode_main_tm_def) apply(rename_tac stp ln rn) apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, rule_tac x = rn in exI, simp) apply(simp add: replicate_Suc[THEN sym] replicate_add[THEN sym] del: replicate_Suc) done from this obtain stpb lnb rnb where stp2: "steps0 (twice_tm_len + 14, Bk # Bk # Bk\(lna) @ Oc # ires, Oc\(Suc (rs + 1)) @ Bk\(rna)) wcode_main_tm stpb = (twice_tm_len + 14 + fourtimes_tm_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rnb))" by blast have "\stp ln rn. steps0 (twice_tm_len + 14 + fourtimes_tm_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rnb)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ Oc # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rn))" apply(rule wcode_jump2) done from this obtain stpc lnc rnc where stp3: "steps0 (twice_tm_len + 14 + fourtimes_tm_len, Bk # Bk # Bk\(lnb) @ Oc # ires, Oc\(Suc (4*rs + 4)) @ Bk\(rnb)) wcode_main_tm stpc = (Suc 0, Bk # Bk\(lnc) @ Oc # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rnc))" by blast from stp1 stp2 stp3 show "?thesis" apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI, rule_tac x = rnc in exI) apply(simp) done qed fun wcode_on_left_moving_3_B :: "bin_inv_t" where "wcode_on_left_moving_3_B ires rs (l, r) = (\ ml mr rn. l = Bk\(ml) @ Oc # Bk # Bk # ires \ r = Bk\(mr) @ Oc\(Suc rs) @ Bk\(rn) \ ml + mr > Suc 0 \ mr > 0 )" fun wcode_on_left_moving_3_O :: "bin_inv_t" where "wcode_on_left_moving_3_O ires rs (l, r) = (\ ln rn. l = Bk # Bk # ires \ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_on_left_moving_3 :: "bin_inv_t" where "wcode_on_left_moving_3 ires rs (l, r) = (wcode_on_left_moving_3_B ires rs (l, r) \ wcode_on_left_moving_3_O ires rs (l, r))" fun wcode_on_checking_3 :: "bin_inv_t" where "wcode_on_checking_3 ires rs (l, r) = (\ ln rn. l = Bk # ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_goon_checking_3 :: "bin_inv_t" where "wcode_goon_checking_3 ires rs (l, r) = (\ ln rn. l = ires \ r = Bk # Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_stop :: "bin_inv_t" where "wcode_stop ires rs (l, r) = (\ ln rn. l = Bk # ires \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" fun wcode_halt_case_inv :: "nat \ bin_inv_t" where "wcode_halt_case_inv st ires rs (l, r) = (if st = 0 then wcode_stop ires rs (l, r) else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r) else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r) else if st = 7 then wcode_goon_checking_3 ires rs (l, r) else False)" fun wcode_halt_case_state :: "config \ nat" where "wcode_halt_case_state (st, l, r) = (if st = 1 then 5 else if st = Suc (Suc 0) then 4 else if st = 7 then 3 else 0)" fun wcode_halt_case_step :: "config \ nat" where "wcode_halt_case_step (st, l, r) = (if st = 1 then length l else 0)" fun wcode_halt_case_measure :: "config \ nat \ nat" where "wcode_halt_case_measure (st, l, r) = (wcode_halt_case_state (st, l, r), wcode_halt_case_step (st, l, r))" definition wcode_halt_case_le :: "(config \ config) set" where "wcode_halt_case_le \ (inv_image lex_pair wcode_halt_case_measure)" lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le" by(auto simp: wcode_halt_case_le_def) declare wcode_on_left_moving_3_B.simps[simp del] wcode_on_left_moving_3_O.simps[simp del] wcode_on_checking_3.simps[simp del] wcode_goon_checking_3.simps[simp del] wcode_on_left_moving_3.simps[simp del] wcode_stop.simps[simp del] lemmas wcode_halt_invs = wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps wcode_on_checking_3.simps wcode_goon_checking_3.simps wcode_on_left_moving_3.simps wcode_stop.simps lemma wcode_on_left_moving_3_mv_Bk[simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) \ wcode_on_left_moving_3 ires rs (tl b, hd b # Bk # list)" apply(simp only: wcode_halt_invs) apply(erule_tac disjE) apply(erule_tac exE)+ apply(rename_tac ml mr rn) apply(case_tac ml, simp) apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI) apply(case_tac mr, force, simp add: exp_ind del: replicate_Suc) apply(case_tac "mr - 1", force, simp add: exp_ind del: replicate_Suc) apply force apply force done lemma wcode_goon_checking_3_cases[simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \ (b = [] \ wcode_stop ires rs ([Bk], list)) \ (b \ [] \ wcode_stop ires rs (Bk # b, list))" apply(auto simp: wcode_halt_invs) done lemma wcode_on_checking_3_mv_Oc[simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \ wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)" by(simp add:wcode_halt_invs) lemma wcode_3_nonempty[simp]: "wcode_on_left_moving_3 ires rs (b, []) = False" "wcode_on_checking_3 ires rs (b, []) = False" "wcode_goon_checking_3 ires rs (b, []) = False" "wcode_on_left_moving_3 ires rs (b, Oc # list) \ b \ []" "wcode_on_checking_3 ires rs (b, Oc # list) = False" "wcode_on_left_moving_3 ires rs (b, Bk # list) \ b \ []" "wcode_on_checking_3 ires rs (b, Bk # list) \ b \ []" "wcode_goon_checking_3 ires rs (b, Oc # list) = False" by(auto simp: wcode_halt_invs) lemma wcode_goon_checking_3_mv_Bk[simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \ wcode_goon_checking_3 ires rs (tl b, hd b # Bk # list)" apply(auto simp: wcode_halt_invs) done lemma t_halt_case_correctness: shows "let P = (\ (st, l, r). st = 0) in let Q = (\ (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in let f = (\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp) in \ n .P (f n) \ Q (f (n::nat))" proof - let ?P = "(\ (st, l, r). st = 0)" let ?Q = "(\ (st, l, r). wcode_halt_case_inv st ires rs (l, r))" let ?f = "(\ stp. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp)" have "\ n. ?P (?f n) \ ?Q (?f (n::nat))" proof(rule_tac halt_lemma2) show "wf wcode_halt_case_le" by auto next { fix na obtain a b c where abc:"?f na = (a,b,c)" by(cases "?f na",auto) hence "\ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_halt_case_le" apply(simp add: step.simps) apply(cases c;cases "hd c") apply(auto simp: wcode_halt_case_le_def lex_pair_def split: if_splits) done } thus "\ na. \ ?P (?f na) \ ?Q (?f na) \ ?Q (?f (Suc na)) \ (?f (Suc na), ?f na) \ wcode_halt_case_le" by blast next show "?Q (?f 0)" apply(simp add: steps.simps wcode_halt_invs) apply(rule_tac x = "Suc m" in exI, simp) apply(rule_tac x = "Suc 0" in exI, auto) done next show "\ ?P (?f 0)" apply(simp add: steps.simps) done qed thus "?thesis" apply(auto) done qed declare wcode_halt_case_inv.simps[simp del] lemma leading_Oc[intro]: "\ xs. ( :: cell list) = Oc # xs" apply(case_tac "rev list", simp) apply(simp add: tape_of_nl_cons) done lemma wcode_halt_case: "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(Suc rs) @ Bk\(rn))" proof - let ?P = "\(st, l, r). st = 0" let ?Q = "\(st, l, r). wcode_halt_case_inv st ires rs (l, r)" let ?f = "steps0 (Suc 0, Bk # Bk \ m @ Oc # Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) wcode_main_tm" from t_halt_case_correctness[of ires rs m n] obtain n where "?P (?f n) \ ?Q (?f n)" by metis thus ?thesis apply(simp add: wcode_halt_case_inv.simps wcode_stop.simps) apply(case_tac "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm n") apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps) by auto qed lemma bl_bin_one[simp]: "bl_bin [Oc] = 1" apply(simp add: bl_bin.simps) done lemma twice_power[intro]: "2 * 2 ^ a = Suc (Suc (2 * bl_bin (Oc \ a)))" apply(induct a, auto simp: bl_bin.simps) done declare replicate_Suc[simp del] lemma wcode_main_tm_lemma_pre: "\args \ []; lm = \ \ \ stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2^(length lm - 1) ) @ Bk\(rn))" proof(induct "length args" arbitrary: args lm rs m n, simp) fix x args lm rs m n assume ind: "\args lm rs m n. \x = length args; (args::nat list) \ []; lm = \ \ \stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\(rn))" and h: "Suc x = length args" "(args::nat list) \ []" "lm = " from h have "\ (a::nat) xs. args = xs @ [a]" apply(rule_tac x = "last args" in exI) apply(rule_tac x = "butlast args" in exI, auto) done from this obtain a xs where "args = xs @ [a]" by blast from h and this show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\(rn))" proof(case_tac "xs::nat list", simp) show "\stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc \ Suc a @ Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin (Oc \ Suc a) + rs * 2 ^ a) @ Bk \ rn)" proof(induct "a" arbitrary: m n rs ires, simp) fix m n rs ires show "\stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc # Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ Suc rs @ Bk \ rn)" apply(rule_tac wcode_halt_case) done next fix a m n rs ires assume ind2: "\m n rs ires. \stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc \ Suc a @ Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin (Oc \ Suc a) + rs * 2 ^ a) @ Bk \ rn)" show " \stp ln rn. steps0 (Suc 0, Bk # Bk \ m @ Oc \ Suc (Suc a) @ Bk # Bk # ires, Bk # Oc \ Suc rs @ Bk \ n) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin (Oc \ Suc (Suc a)) + rs * 2 ^ Suc a) @ Bk \ rn)" proof - have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rn))" apply(simp add: tape_of_nat) using wcode_double_case[of m "Oc\(a) @ Bk # Bk # ires" rs n] apply(simp add: replicate_Suc) done from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stpa = (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rna))" by blast moreover have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rna)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin () + (2*rs + 2) * 2 ^ a) @ Bk\(rn))" using ind2[of lna ires "2*rs + 2" rna] by(simp add: tape_of_list_def tape_of_nat_def) from this obtain stpb lnb rnb where stp2: "steps0 (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (2 * rs + 2)) @ Bk\(rna)) wcode_main_tm stpb = (0, Bk # ires, Bk # Oc # Bk\(lnb) @ Bk # Bk # Oc\(bl_bin () + (2*rs + 2) * 2 ^ a) @ Bk\(rnb))" by blast from stp1 and stp2 show "?thesis" apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: tape_of_nat_def) apply(simp add: bl_bin.simps replicate_Suc) apply(auto) done qed qed next fix aa list assume g: "Suc x = length args" "args \ []" "lm = " "args = xs @ [a::nat]" "xs = (aa::nat) # list" thus "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\(rn))" proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev del: subst_all, simp only: tape_of_nl_cons_app1, simp del: subst_all) fix m n rs args lm have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rn))" proof(simp add: tape_of_nl_rev) have "\ xs. () = Oc # xs" by auto from this obtain xs where "() = Oc # xs" .. thus "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ @ Bk # Bk # ires, Bk # Oc\(5 + 4 * rs) @ Bk\(rn))" apply(simp) using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n] apply(simp) done qed from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # rev () @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stpa = (Suc 0, Bk # Bk\(lna) @ rev () @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rna))" by blast from g have "\ stp ln rn. steps0 (Suc 0, Bk # Bk\(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rna)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()+ (4*rs + 4) * 2^(length () - 1) ) @ Bk\(rn))" apply(rule_tac args = "(aa::nat)#list" in ind, simp_all) done from this obtain stpb lnb rnb where stp2: "steps0 (Suc 0, Bk # Bk\(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, Bk # Oc\(Suc (4*rs + 4)) @ Bk\(rna)) wcode_main_tm stpb = (0, Bk # ires, Bk # Oc # Bk\(lnb) @ Bk # Bk # Oc\(bl_bin ()+ (4*rs + 4) * 2^(length () - 1) ) @ Bk\(rnb))" by blast from stp1 and stp2 and h show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc # Bk # @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin (Oc\(Suc aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))) @ Bk\(rn))" apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: tape_of_nl_rev) done next fix ab m n rs args lm assume ind2: "\ m n rs args lm. \lm = ; args = aa # list @ [ab]\ \ \stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin () + rs * 2 ^ (length () - Suc 0)) @ Bk\(rn))" and k: "args = aa # list @ [Suc ab]" "lm = " show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires,Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin () + rs * 2 ^ (length () - Suc 0)) @ Bk\(rn))" proof(simp add: tape_of_nl_cons_app1) have "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc\(Suc (Suc ab)) @ Bk # @ Bk # Bk # ires, Bk # Oc # Oc\(rs) @ Bk\(n)) wcode_main_tm stp = (Suc 0, Bk # Bk\(ln) @ Oc\(Suc ab) @ Bk # @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rn))" using wcode_double_case[of m "Oc\(ab) @ Bk # @ Bk # Bk # ires" rs n] apply(simp add: replicate_Suc) done from this obtain stpa lna rna where stp1: "steps0 (Suc 0, Bk # Bk\(m) @ Oc\(Suc (Suc ab)) @ Bk # @ Bk # Bk # ires, Bk # Oc # Oc\(rs) @ Bk\(n)) wcode_main_tm stpa = (Suc 0, Bk # Bk\(lna) @ Oc\(Suc ab) @ Bk # @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rna))" by blast from k have "\ stp ln rn. steps0 (Suc 0, Bk # Bk\(lna) @ @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rna)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ( ) + (2*rs + 2)* 2^(length () - Suc 0)) @ Bk\(rn))" apply(rule_tac ind2, simp_all) done from this obtain stpb lnb rnb where stp2: "steps0 (Suc 0, Bk # Bk\(lna) @ @ Bk # Bk # ires, Bk # Oc\(Suc (2*rs + 2)) @ Bk\(rna)) wcode_main_tm stpb = (0, Bk # ires, Bk # Oc # Bk\(lnb) @ Bk # Bk # Oc\(bl_bin ( ) + (2*rs + 2)* 2^(length () - Suc 0)) @ Bk\(rnb))" by blast from stp1 and stp2 show "\stp ln rn. steps0 (Suc 0, Bk # Bk\(m) @ Oc\(Suc (Suc ab)) @ Bk # @ Bk # Bk # ires, Bk # Oc\(Suc rs) @ Bk\(n)) wcode_main_tm stp = (0, Bk # ires, Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin (Oc\(Suc aa) @ Bk # ) + rs * (2 * 2 ^ (aa + length ()))) @ Bk\(rn))" apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: tape_of_nl_cons_app1 replicate_Suc) done qed qed qed qed definition wcode_prepare_tm :: "instr list" where "wcode_prepare_tm \ [(WO, 2), (L, 1), (L, 3), (R, 2), (R, 4), (WB, 3), (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5), (WO, 7), (L, 0)]" fun wprepare_add_one :: "nat \ nat list \ tape \ bool" where "wprepare_add_one m lm (l, r) = (\ rn. l = [] \ (r = @ Bk\(rn) \ r = Bk # @ Bk\(rn)))" fun wprepare_goto_first_end :: "nat \ nat list \ tape \ bool" where "wprepare_goto_first_end m lm (l, r) = (\ ml mr rn. l = Oc\(ml) \ r = Oc\(mr) @ Bk # @ Bk\(rn) \ ml + mr = Suc (Suc m))" fun wprepare_erase :: "nat \ nat list \ tape \ bool" where "wprepare_erase m lm (l, r) = (\ rn. l = Oc\(Suc m) \ tl r = Bk # @ Bk\(rn))" fun wprepare_goto_start_pos_B :: "nat \ nat list \ tape \ bool" where "wprepare_goto_start_pos_B m lm (l, r) = (\ rn. l = Bk # Oc\(Suc m) \ r = Bk # @ Bk\(rn))" fun wprepare_goto_start_pos_O :: "nat \ nat list \ tape \ bool" where "wprepare_goto_start_pos_O m lm (l, r) = (\ rn. l = Bk # Bk # Oc\(Suc m) \ r = @ Bk\(rn))" fun wprepare_goto_start_pos :: "nat \ nat list \ tape \ bool" where "wprepare_goto_start_pos m lm (l, r) = (wprepare_goto_start_pos_B m lm (l, r) \ wprepare_goto_start_pos_O m lm (l, r))" fun wprepare_loop_start_on_rightmost :: "nat \ nat list \ tape \ bool" where "wprepare_loop_start_on_rightmost m lm (l, r) = (\ rn mr. rev l @ r = Oc\(Suc m) @ Bk # Bk # @ Bk\(rn) \ l \ [] \ r = Oc\(mr) @ Bk\(rn))" fun wprepare_loop_start_in_middle :: "nat \ nat list \ tape \ bool" where "wprepare_loop_start_in_middle m lm (l, r) = (\ rn (mr:: nat) (lm1::nat list). rev l @ r = Oc\(Suc m) @ Bk # Bk # @ Bk\(rn) \ l \ [] \ r = Oc\(mr) @ Bk # @ Bk\(rn) \ lm1 \ [])" fun wprepare_loop_start :: "nat \ nat list \ tape \ bool" where "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \ wprepare_loop_start_in_middle m lm (l, r))" fun wprepare_loop_goon_on_rightmost :: "nat \ nat list \ tape \ bool" where "wprepare_loop_goon_on_rightmost m lm (l, r) = (\ rn. l = Bk # @ Bk # Bk # Oc\(Suc m) \ r = Bk\(rn))" fun wprepare_loop_goon_in_middle :: "nat \ nat list \ tape \ bool" where "wprepare_loop_goon_in_middle m lm (l, r) = (\ rn (mr:: nat) (lm1::nat list). rev l @ r = Oc\(Suc m) @ Bk # Bk # @ Bk\(rn) \ l \ [] \ (if lm1 = [] then r = Oc\(mr) @ Bk\(rn) else r = Oc\(mr) @ Bk # @ Bk\(rn)) \ mr > 0)" fun wprepare_loop_goon :: "nat \ nat list \ tape \ bool" where "wprepare_loop_goon m lm (l, r) = (wprepare_loop_goon_in_middle m lm (l, r) \ wprepare_loop_goon_on_rightmost m lm (l, r))" fun wprepare_add_one2 :: "nat \ nat list \ tape \ bool" where "wprepare_add_one2 m lm (l, r) = (\ rn. l = Bk # Bk # @ Bk # Bk # Oc\(Suc m) \ (r = [] \ tl r = Bk\(rn)))" fun wprepare_stop :: "nat \ nat list \ tape \ bool" where "wprepare_stop m lm (l, r) = (\ rn. l = Bk # @ Bk # Bk # Oc\(Suc m) \ r = Bk # Oc # Bk\(rn))" fun wprepare_inv :: "nat \ nat \ nat list \ tape \ bool" where "wprepare_inv st m lm (l, r) = (if st = 0 then wprepare_stop m lm (l, r) else if st = Suc 0 then wprepare_add_one m lm (l, r) else if st = Suc (Suc 0) then wprepare_goto_first_end m lm (l, r) else if st = Suc (Suc (Suc 0)) then wprepare_erase m lm (l, r) else if st = 4 then wprepare_goto_start_pos m lm (l, r) else if st = 5 then wprepare_loop_start m lm (l, r) else if st = 6 then wprepare_loop_goon m lm (l, r) else if st = 7 then wprepare_add_one2 m lm (l, r) else False)" fun wprepare_stage :: "config \ nat" where "wprepare_stage (st, l, r) = (if st \ 1 \ st \ 4 then 3 else if st = 5 \ st = 6 then 2 else 1)" fun wprepare_state :: "config \ nat" where "wprepare_state (st, l, r) = (if st = 1 then 4 else if st = Suc (Suc 0) then 3 else if st = Suc (Suc (Suc 0)) then 2 else if st = 4 then 1 else if st = 7 then 2 else 0)" fun wprepare_step :: "config \ nat" where "wprepare_step (st, l, r) = (if st = 1 then (if hd r = Oc then Suc (length l) else 0) else if st = Suc (Suc 0) then length r else if st = Suc (Suc (Suc 0)) then (if hd r = Oc then 1 else 0) else if st = 4 then length r else if st = 5 then Suc (length r) else if st = 6 then (if r = [] then 0 else Suc (length r)) else if st = 7 then (if (r \ [] \ hd r = Oc) then 0 else 1) else 0)" fun wcode_prepare_measure :: "config \ nat \ nat \ nat" where "wcode_prepare_measure (st, l, r) = (wprepare_stage (st, l, r), wprepare_state (st, l, r), wprepare_step (st, l, r))" definition wcode_prepare_le :: "(config \ config) set" where "wcode_prepare_le \ (inv_image lex_triple wcode_prepare_measure)" lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le" by(auto simp: wcode_prepare_le_def lex_triple_def) declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del] wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del] wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del] wprepare_add_one2.simps[simp del] lemmas wprepare_invs = wprepare_add_one.simps wprepare_goto_first_end.simps wprepare_erase.simps wprepare_goto_start_pos.simps wprepare_loop_start.simps wprepare_loop_goon.simps wprepare_add_one2.simps declare wprepare_inv.simps[simp del] lemma fetch_wcode_prepare_tm[simp]: "fetch wcode_prepare_tm (Suc 0) Bk = (WO, 2)" "fetch wcode_prepare_tm (Suc 0) Oc = (L, 1)" "fetch wcode_prepare_tm (Suc (Suc 0)) Bk = (L, 3)" "fetch wcode_prepare_tm (Suc (Suc 0)) Oc = (R, 2)" "fetch wcode_prepare_tm (Suc (Suc (Suc 0))) Bk = (R, 4)" "fetch wcode_prepare_tm (Suc (Suc (Suc 0))) Oc = (WB, 3)" "fetch wcode_prepare_tm 4 Bk = (R, 4)" "fetch wcode_prepare_tm 4 Oc = (R, 5)" "fetch wcode_prepare_tm 5 Oc = (R, 5)" "fetch wcode_prepare_tm 5 Bk = (R, 6)" "fetch wcode_prepare_tm 6 Oc = (R, 5)" "fetch wcode_prepare_tm 6 Bk = (R, 7)" "fetch wcode_prepare_tm 7 Oc = (L, 0)" "fetch wcode_prepare_tm 7 Bk = (WO, 7)" unfolding fetch.simps wcode_prepare_tm_def nth_of.simps numeral_eqs_upto_12 by auto lemma wprepare_add_one_nonempty_snd[simp]: "lm \ [] \ wprepare_add_one m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_goto_first_end_nonempty_snd[simp]: "lm \ [] \ wprepare_goto_first_end m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_erase_nonempty_snd[simp]: "lm \ [] \ wprepare_erase m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_goto_start_pos_nonempty_snd[simp]: "lm \ [] \ wprepare_goto_start_pos m lm (b, []) = False" apply(simp add: wprepare_invs) done lemma wprepare_loop_start_empty_nonempty_fst[simp]: "\lm \ []; wprepare_loop_start m lm (b, [])\ \ b \ []" apply(simp add: wprepare_invs) done lemma rev_eq: "rev xs = rev ys \ xs = ys" by auto lemma wprepare_loop_goon_Bk_empty_snd[simp]: "\lm \ []; wprepare_loop_start m lm (b, [])\ \ wprepare_loop_goon m lm (Bk # b, [])" apply(simp only: wprepare_invs) apply(erule_tac disjE) apply(rule_tac disjI2) apply(simp, auto) apply(rule_tac rev_eq, simp add: tape_of_nl_rev) done lemma wprepare_loop_goon_nonempty_fst[simp]: "\lm \ []; wprepare_loop_goon m lm (b, [])\ \ b \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_add_one2_Bk_empty[simp]:"\lm \ []; wprepare_loop_goon m lm (b, [])\ \ wprepare_add_one2 m lm (Bk # b, [])" apply(simp only: wprepare_invs, auto split: if_splits) done lemma wprepare_add_one2_nonempty_fst[simp]: "wprepare_add_one2 m lm (b, []) \ b \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_add_one2_Oc[simp]: "wprepare_add_one2 m lm (b, []) \ wprepare_add_one2 m lm (b, [Oc])" apply(simp only: wprepare_invs, auto) done lemma Bk_not_tape_start[simp]: "(Bk # list = <(m::nat) # lm> @ ys) = False" apply(case_tac lm, auto simp: tape_of_nl_cons replicate_Suc) done lemma wprepare_goto_first_end_cases[simp]: "\lm \ []; wprepare_add_one m lm (b, Bk # list)\ \ (b = [] \ wprepare_goto_first_end m lm ([], Oc # list)) \ (b \ [] \ wprepare_goto_first_end m lm (b, Oc # list))" apply(simp only: wprepare_invs) apply(auto simp: tape_of_nl_cons split: if_splits) apply(cases lm, auto simp add:tape_of_list_def replicate_Suc) done lemma wprepare_goto_first_end_Bk_nonempty_fst[simp]: "wprepare_goto_first_end m lm (b, Bk # list) \ b \ []" apply(simp only: wprepare_invs , auto simp: replicate_Suc) done declare replicate_Suc[simp] lemma wprepare_erase_elem_Bk_rest[simp]: "wprepare_goto_first_end m lm (b, Bk # list) \ wprepare_erase m lm (tl b, hd b # Bk # list)" by(simp add: wprepare_invs) lemma wprepare_erase_Bk_nonempty_fst[simp]: "wprepare_erase m lm (b, Bk # list) \ b \ []" by(simp add: wprepare_invs) lemma wprepare_goto_start_pos_Bk[simp]: "wprepare_erase m lm (b, Bk # list) \ wprepare_goto_start_pos m lm (Bk # b, list)" apply(simp only: wprepare_invs, auto) done lemma wprepare_add_one_Bk_nonempty_snd[simp]: "\wprepare_add_one m lm (b, Bk # list)\ \ list \ []" apply(simp only: wprepare_invs) apply(case_tac lm, simp_all add: tape_of_list_def tape_of_nat_def, auto) done lemma wprepare_goto_first_end_nonempty_snd_tl[simp]: "\lm \ []; wprepare_goto_first_end m lm (b, Bk # list)\ \ list \ []" by(simp only: wprepare_invs, auto) lemma wprepare_erase_Bk_nonempty_list[simp]: "\lm \ []; wprepare_erase m lm (b, Bk # list)\ \ list \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_goto_start_pos_Bk_nonempty[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ list \ []" by(cases lm;cases list;simp only: wprepare_invs, auto) lemma wprepare_goto_start_pos_Bk_nonempty_fst[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ b \ []" apply(simp only: wprepare_invs) apply(auto) done lemma wprepare_loop_goon_Bk_nonempty[simp]: "\lm \ []; wprepare_loop_goon m lm (b, Bk # list)\ \ b \ []" apply(simp only: wprepare_invs, auto) done lemma wprepare_loop_goon_wprepare_add_one2_cases[simp]: "\lm \ []; wprepare_loop_goon m lm (b, Bk # list)\ \ (list = [] \ wprepare_add_one2 m lm (Bk # b, [])) \ (list \ [] \ wprepare_add_one2 m lm (Bk # b, list))" unfolding wprepare_invs apply(cases list;auto split:nat.split if_splits) by (metis list.sel(3) tl_replicate) lemma wprepare_add_one2_nonempty[simp]: "wprepare_add_one2 m lm (b, Bk # list) \ b \ []" apply(simp only: wprepare_invs, simp) done lemma wprepare_add_one2_cases[simp]: "wprepare_add_one2 m lm (b, Bk # list) \ (list = [] \ wprepare_add_one2 m lm (b, [Oc])) \ (list \ [] \ wprepare_add_one2 m lm (b, Oc # list))" apply(simp only: wprepare_invs, auto) done lemma wprepare_goto_first_end_cases_Oc[simp]: "wprepare_goto_first_end m lm (b, Oc # list) \ (b = [] \ wprepare_goto_first_end m lm ([Oc], list)) \ (b \ [] \ wprepare_goto_first_end m lm (Oc # b, list))" apply(simp only: wprepare_invs, auto) apply(rule_tac x = 1 in exI, auto) apply(rename_tac ml mr rn) apply(case_tac mr, simp_all add: ) apply(case_tac ml, simp_all add: ) apply(rule_tac x = "Suc ml" in exI, simp_all add: ) apply(rule_tac x = "mr - 1" in exI, simp) done lemma wprepare_erase_nonempty[simp]: "wprepare_erase m lm (b, Oc # list) \ b \ []" apply(simp only: wprepare_invs, auto simp: ) done lemma wprepare_erase_Bk[simp]: "wprepare_erase m lm (b, Oc # list) \ wprepare_erase m lm (b, Bk # list)" apply(simp only:wprepare_invs, auto simp: ) done lemma wprepare_goto_start_pos_Bk_move[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Bk # list)\ \ wprepare_goto_start_pos m lm (Bk # b, list)" apply(simp only:wprepare_invs, auto) apply(case_tac [!] lm, simp, simp_all) done lemma wprepare_loop_start_b_nonempty[simp]: "wprepare_loop_start m lm (b, aa) \ b \ []" apply(simp only:wprepare_invs, auto) done lemma exists_exp_of_Bk[elim]: "Bk # list = Oc\(mr) @ Bk\(rn) \ \rn. list = Bk\(rn)" apply(case_tac mr, simp_all) apply(case_tac rn, simp_all) done lemma wprepare_loop_start_in_middle_Bk_False[simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False" by(auto) declare wprepare_loop_start_in_middle.simps[simp del] declare wprepare_loop_start_on_rightmost.simps[simp del] wprepare_loop_goon_in_middle.simps[simp del] wprepare_loop_goon_on_rightmost.simps[simp del] lemma wprepare_loop_goon_in_middle_Bk_False[simp]: "wprepare_loop_goon_in_middle m lm (Bk # b, []) = False" apply(simp add: wprepare_loop_goon_in_middle.simps, auto) done lemma wprepare_loop_goon_Bk[simp]: "\lm \ []; wprepare_loop_start m lm (b, [Bk])\ \ wprepare_loop_goon m lm (Bk # b, [])" unfolding wprepare_invs apply(auto simp add: wprepare_loop_goon_on_rightmost.simps wprepare_loop_start_on_rightmost.simps) apply(rule_tac rev_eq) apply(simp add: tape_of_nl_rev) apply(simp add: exp_ind replicate_Suc[THEN sym] del: replicate_Suc) done lemma wprepare_loop_goon_in_middle_Bk_False2[simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista) \ wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False" apply(auto simp: wprepare_loop_start_on_rightmost.simps wprepare_loop_goon_in_middle.simps) done lemma wprepare_loop_goon_on_rightbmost_Bk_False[simp]: "\lm \ []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\ \ wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)" apply(simp only: wprepare_loop_start_on_rightmost.simps wprepare_loop_goon_on_rightmost.simps, auto simp: tape_of_nl_rev) apply(simp add: replicate_Suc[THEN sym] exp_ind tape_of_nl_rev del: replicate_Suc) by (meson Cons_replicate_eq) lemma wprepare_loop_goon_in_middle_Bk_False3[simp]: assumes "lm \ []" "wprepare_loop_start_in_middle m lm (b, Bk # a # lista)" shows "wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)" (is "?t1") and "wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False" (is ?t2) proof - from assms obtain rn mr lm1 where *:"rev b @ Oc \ mr @ Bk # = Oc # Oc \ m @ Bk # Bk # " "b \ []" "Bk # a # lista = Oc \ mr @ Bk # @ Bk \ rn" "lm1 \ []" by(auto simp add: wprepare_loop_start_in_middle.simps) thus ?t1 apply(simp add: wprepare_loop_start_in_middle.simps wprepare_loop_goon_in_middle.simps, auto) apply(rule_tac x = rn in exI, simp) apply(case_tac mr, simp_all add: ) apply(case_tac lm1, simp) apply(rule_tac x = "Suc (hd lm1)" in exI, simp) apply(rule_tac x = "tl lm1" in exI) apply(case_tac "tl lm1", simp_all add: tape_of_list_def tape_of_nat_def) done from * show ?t2 apply(simp add: wprepare_loop_start_in_middle.simps wprepare_loop_goon_on_rightmost.simps del:split_head_repeat, auto simp del:split_head_repeat) apply(case_tac mr) apply(case_tac "lm1::nat list", simp_all, case_tac "tl lm1", simp_all) apply(auto simp add: tape_of_list_def ) apply(case_tac [!] rna, simp_all add: ) apply(case_tac mr, simp_all add: ) apply(case_tac lm1, simp, case_tac list, simp) apply(simp_all add: tape_of_nat_def) by (metis Bk_not_tape_start tape_of_list_def tape_of_nat_list.elims) qed lemma wprepare_loop_goon_Bk2[simp]: "\lm \ []; wprepare_loop_start m lm (b, Bk # a # lista)\ \ wprepare_loop_goon m lm (Bk # b, a # lista)" apply(simp add: wprepare_loop_start.simps wprepare_loop_goon.simps) apply(erule_tac disjE, simp, auto) done lemma start_2_goon: "\lm \ []; wprepare_loop_start m lm (b, Bk # list)\ \ (list = [] \ wprepare_loop_goon m lm (Bk # b, [])) \ (list \ [] \ wprepare_loop_goon m lm (Bk # b, list))" apply(case_tac list, auto) done lemma add_one_2_add_one: "wprepare_add_one m lm (b, Oc # list) \ (hd b = Oc \ (b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)) \ (b \ [] \ wprepare_add_one m lm (tl b, Oc # Oc # list))) \ (hd b \ Oc \ (b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)) \ (b \ [] \ wprepare_add_one m lm (tl b, hd b # Oc # list)))" unfolding wprepare_add_one.simps by auto lemma wprepare_loop_start_on_rightmost_Oc[simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \ wprepare_loop_start_on_rightmost m lm (Oc # b, list)" apply(simp add: wprepare_loop_start_on_rightmost.simps) by (metis Cons_replicate_eq cell.distinct(1) list.sel(3) self_append_conv2 tl_append2 tl_replicate) lemma wprepare_loop_start_in_middle_Oc[simp]: assumes "wprepare_loop_start_in_middle m lm (b, Oc # list)" shows "wprepare_loop_start_in_middle m lm (Oc # b, list)" proof - from assms obtain mr lm1 rn where "rev b @ Oc \ mr @ Bk # = Oc # Oc \ m @ Bk # Bk # " "Oc # list = Oc \ mr @ Bk # @ Bk \ rn" "lm1 \ []" by(auto simp add: wprepare_loop_start_in_middle.simps) thus ?thesis apply(auto simp add: wprepare_loop_start_in_middle.simps) apply(rule_tac x = rn in exI, auto) apply(case_tac mr, simp, simp add: ) apply(rule_tac x = "mr - 1" in exI, simp) apply(rule_tac x = lm1 in exI, simp) done qed lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \ wprepare_loop_start m lm (Oc # b, list)" apply(simp add: wprepare_loop_start.simps) apply(erule_tac disjE, simp_all ) done lemma wprepare_loop_goon_Oc_nonempty[simp]: "wprepare_loop_goon m lm (b, Oc # list) \ b \ []" apply(simp add: wprepare_loop_goon.simps wprepare_loop_goon_in_middle.simps wprepare_loop_goon_on_rightmost.simps) apply(auto) done lemma wprepare_goto_start_pos_Oc_nonempty[simp]: "wprepare_goto_start_pos m lm (b, Oc # list) \ b \ []" apply(simp add: wprepare_goto_start_pos.simps) done lemma wprepare_loop_goon_on_rightmost_Oc_False[simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False" apply(simp add: wprepare_loop_goon_on_rightmost.simps) done lemma wprepare_loop1: "\rev b @ Oc\(mr) = Oc\(Suc m) @ Bk # Bk # ; b \ []; 0 < mr; Oc # list = Oc\(mr) @ Bk\(rn)\ \ wprepare_loop_start_on_rightmost m lm (Oc # b, list)" apply(simp add: wprepare_loop_start_on_rightmost.simps) apply(rule_tac x = rn in exI, simp) apply(case_tac mr, simp, simp) done lemma wprepare_loop2: "\rev b @ Oc\(mr) @ Bk # = Oc\(Suc m) @ Bk # Bk # ; b \ []; Oc # list = Oc\(mr) @ Bk # <(a::nat) # lista> @ Bk\(rn)\ \ wprepare_loop_start_in_middle m lm (Oc # b, list)" apply(simp add: wprepare_loop_start_in_middle.simps) apply(rule_tac x = rn in exI, simp) apply(case_tac mr, simp_all add: ) apply(rename_tac nat) apply(rule_tac x = nat in exI, simp) apply(rule_tac x = "a#lista" in exI, simp) done lemma wprepare_loop_goon_in_middle_cases[simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \ wprepare_loop_start_on_rightmost m lm (Oc # b, list) \ wprepare_loop_start_in_middle m lm (Oc # b, list)" apply(simp add: wprepare_loop_goon_in_middle.simps split: if_splits) apply(rename_tac lm1) apply(case_tac lm1, simp_all add: wprepare_loop1 wprepare_loop2) done lemma wprepare_add_one_b[simp]: "wprepare_add_one m lm (b, Oc # list) \ b = [] \ wprepare_add_one m lm ([], Bk # Oc # list)" "wprepare_loop_goon m lm (b, Oc # list) \ wprepare_loop_start m lm (Oc # b, list)" apply(auto simp add: wprepare_add_one.simps wprepare_loop_goon.simps wprepare_loop_start.simps) done lemma wprepare_loop_start_on_rightmost_Oc2[simp]: "wprepare_goto_start_pos m [a] (b, Oc # list) \ wprepare_loop_start_on_rightmost m [a] (Oc # b, list) " apply(auto simp: wprepare_goto_start_pos.simps wprepare_loop_start_on_rightmost.simps) apply(rename_tac rn) apply(rule_tac x = rn in exI, simp) apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc) done lemma wprepare_loop_start_in_middle_2_Oc[simp]: "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list) \wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)" apply(auto simp: wprepare_goto_start_pos.simps wprepare_loop_start_in_middle.simps) apply(rename_tac rn) apply(rule_tac x = rn in exI, simp) apply(simp add: exp_ind[THEN sym]) apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp) apply(simp add: tape_of_nl_cons) done lemma wprepare_loop_start_Oc2[simp]: "\lm \ []; wprepare_goto_start_pos m lm (b, Oc # list)\ \ wprepare_loop_start m lm (Oc # b, list)" by(cases lm;cases "tl lm", auto simp add: wprepare_loop_start.simps) lemma wprepare_add_one2_Oc_nonempty[simp]: "wprepare_add_one2 m lm (b, Oc # list) \ b \ []" apply(auto simp: wprepare_add_one2.simps) done lemma add_one_2_stop: "wprepare_add_one2 m lm (b, Oc # list) \ wprepare_stop m lm (tl b, hd b # Oc # list)" apply(simp add: wprepare_add_one2.simps) done declare wprepare_stop.simps[simp del] lemma wprepare_correctness: assumes h: "lm \ []" shows "let P = (\ (st, l, r). st = 0) in let Q = (\ (st, l, r). wprepare_inv st m lm (l, r)) in let f = (\ stp. steps0 (Suc 0, [], ()) wcode_prepare_tm stp) in \ n .P (f n) \ Q (f n)" proof - let ?P = "(\ (st, l, r). st = 0)" let ?Q = "(\ (st, l, r). wprepare_inv st m lm (l, r))" let ?f = "(\ stp. steps0 (Suc 0, [], ()) wcode_prepare_tm stp)" have "\ n. ?P (?f n) \ ?Q (?f n)" proof(rule_tac halt_lemma2) show "\ n. \ ?P (?f n) \ ?Q (?f n) \ ?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wcode_prepare_le" using h apply(rule_tac allI, rule_tac impI) apply(rename_tac n) apply(case_tac "?f n", simp add: step.simps) apply(rename_tac c) apply(case_tac c, simp, case_tac [2] aa) apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def lex_triple_def lex_pair_def split: if_splits) (* slow *) apply(simp_all add: start_2_goon start_2_start add_one_2_add_one add_one_2_stop) apply(auto simp: wprepare_add_one2.simps) done qed (auto simp add: steps.simps wprepare_inv.simps wprepare_invs) thus "?thesis" apply(auto) done qed lemma composable_tm_wcode_prepare_tm[intro]: "composable_tm (wcode_prepare_tm, 0)" apply(simp add:composable_tm.simps wcode_prepare_tm_def) done lemma is_28_even[intro]: "(28 + (length twice_compile_tm + length fourtimes_compile_tm)) mod 2 = 0" by(auto simp: twice_compile_tm_def fourtimes_compile_tm_def) lemma b_le_28[elim]: "(a, b) \ set wcode_main_first_part_tm \ b \ (28 + (length twice_compile_tm + length fourtimes_compile_tm)) div 2" apply(auto simp: wcode_main_first_part_tm_def twice_tm_def) done lemma composable_tm_change_termi: assumes "composable_tm (tp, 0)" shows "list_all (\(acn, st). (st \ Suc (length tp div 2))) (adjust0 tp)" proof - { fix acn st n assume "tp ! n = (acn, st)" "n < length tp" "0 < st" hence "(acn, st)\set tp" by (metis nth_mem) with assms composable_tm.simps have "st \ length tp div 2 + 0" by auto hence "st \ Suc (length tp div 2)" by auto } thus ?thesis by(auto simp: composable_tm.simps List.list_all_length adjust.simps split: if_splits prod.split) qed lemma composable_tm_shift: assumes "list_all (\(acn, st). (st \ y)) tp" shows "list_all (\(acn, st). (st \ y + off)) (shift tp off)" proof - have [dest!]:"\ P Q n. \n. Q n \ P n \ Q n \ P n" by metis from assms show ?thesis by(auto simp: composable_tm.simps List.list_all_length shift.simps) qed declare length_tp'[simp del] lemma length_mopup_1[simp]: "length (mopup_n_tm (Suc 0)) = 16" apply(auto simp: mopup_n_tm.simps) done lemma twice_plus_28_elim[elim]: "(a, b) \ set (shift (adjust0 twice_compile_tm) 12) \ b \ (28 + (length twice_compile_tm + length fourtimes_compile_tm)) div 2" apply(simp add: twice_compile_tm_def fourtimes_compile_tm_def) proof - assume g: "(a, b) \ set (shift (adjust (tm_of abc_twice @ shift (mopup_n_tm (Suc 0)) (length (tm_of abc_twice) div 2)) (Suc ((length (tm_of abc_twice) + 16) div 2))) 12)" moreover have "length (tm_of abc_twice) mod 2 = 0" by auto moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto ultimately have "list_all (\(acn, st). (st \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) (shift (adjust0 twice_compile_tm) 12)" proof(auto simp add: mod_ex1) assume "even (length (tm_of abc_twice))" then obtain q where q:"length (tm_of abc_twice) = 2 * q" by auto assume "even (length (tm_of abc_fourtimes))" then obtain qa where qa:"length (tm_of abc_fourtimes) = 2 * qa" by auto note h = q qa hence "list_all (\(acn, st). st \ (18 + (q + qa)) + 12) (shift (adjust0 twice_compile_tm) 12)" proof(rule_tac composable_tm_shift twice_compile_tm_def) have "list_all (\(acn, st). st \ Suc (length twice_compile_tm div 2)) (adjust0 twice_compile_tm)" by(rule_tac composable_tm_change_termi, auto) thus "list_all (\(acn, st). st \ 18 + (q + qa)) (adjust0 twice_compile_tm)" using h apply(simp add: twice_compile_tm_def, auto simp: List.list_all_length) done qed thus "list_all (\(acn, st). st \ 30 + (length (tm_of abc_twice) div 2 + length (tm_of abc_fourtimes) div 2)) (shift (adjust0 twice_compile_tm) 12)" using h by simp qed thus "b \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2" using g apply(auto simp:twice_compile_tm_def) apply(simp add: Ball_set[THEN sym]) apply(erule_tac x = "(a, b)" in ballE, simp, simp) done qed lemma length_plus_28_elim2[elim]: "(a, b) \ set (shift (adjust0 fourtimes_compile_tm) (twice_tm_len + 13)) \ b \ (28 + (length twice_compile_tm + length fourtimes_compile_tm)) div 2" apply(simp add: twice_compile_tm_def fourtimes_compile_tm_def twice_tm_len_def) proof - assume g: "(a, b) \ set (shift (adjust (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) (length (tm_of abc_fourtimes) div 2)) (Suc ((length (tm_of abc_fourtimes) + 16) div 2))) (length twice_tm div 2 + 13))" moreover have "length (tm_of abc_twice) mod 2 = 0" by auto moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto ultimately have "list_all (\(acn, st). (st \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) (shift (adjust0 (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) (length (tm_of abc_fourtimes) div 2))) (length twice_tm div 2 + 13))" proof(auto simp: mod_ex1 twice_tm_def twice_compile_tm_def) assume "even (length (tm_of abc_twice))" then obtain q where q:"length (tm_of abc_twice) = 2 * q" by auto assume "even (length (tm_of abc_fourtimes))" then obtain qa where qa:"length (tm_of abc_fourtimes) = 2 * qa" by auto note h = q qa hence "list_all (\(acn, st). st \ (9 + qa + (21 + q))) (shift (adjust0 (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) qa)) (21 + q))" proof(rule_tac composable_tm_shift twice_compile_tm_def) have "list_all (\(acn, st). st \ Suc (length (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) qa) div 2)) (adjust0 (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) qa))" apply(rule_tac composable_tm_change_termi) using composable_fourtimes_tm h apply(simp add: fourtimes_compile_tm_def) done thus "list_all (\(acn, st). st \ 9 + qa) (adjust (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) qa) (Suc (length (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) qa) div 2)))" using h apply(simp) done qed thus "list_all (\(acn, st). st \ 30 + (length (tm_of abc_twice) div 2 + length (tm_of abc_fourtimes) div 2)) (shift (adjust (tm_of abc_fourtimes @ shift (mopup_n_tm (Suc 0)) (length (tm_of abc_fourtimes) div 2)) (9 + length (tm_of abc_fourtimes) div 2)) (21 + length (tm_of abc_twice) div 2))" apply(subgoal_tac "qa + q = q + qa") apply(simp add: h) apply(simp) done qed thus "b \ (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2" using g apply(simp add: Ball_set[THEN sym]) apply(erule_tac x = "(a, b)" in ballE, simp, simp) done qed lemma composable_tm_wcode_main_tm[intro]: "composable_tm (wcode_main_tm, 0)" by(auto simp: wcode_main_tm_def composable_tm.simps twice_tm_def fourtimes_tm_def ) lemma prepare_mainpart_lemma: "args \ [] \ \ stp ln rn. steps0 (Suc 0, [], ) (wcode_prepare_tm |+| wcode_main_tm) stp = (0, Bk # Oc\(Suc m), Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()) @ Bk\(rn))" proof - let ?P1 = "(\ (l, r). (l::cell list) = [] \ r = )" let ?Q1 = "(\ (l, r). wprepare_stop m args (l, r))" let ?P2 = ?Q1 let ?Q2 = "(\ (l, r). (\ ln rn. l = Bk # Oc\(Suc m) \ r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()) @ Bk\(rn)))" let ?P3 = "\ tp. False" assume h: "args \ []" have " \?P1\ wcode_prepare_tm |+| wcode_main_tm \?Q2\" proof(rule_tac Hoare_plus_halt) show " \?P1\ wcode_prepare_tm \?Q1\ " proof(rule_tac Hoare_haltI, auto) show "\n. is_final (steps0 (Suc 0, [], ) wcode_prepare_tm n) \ wprepare_stop m args holds_for steps0 (Suc 0, [], ) wcode_prepare_tm n" using wprepare_correctness[of args m,OF h] apply(auto simp add: wprepare_inv.simps) by (metis holds_for.simps is_finalI) qed next show " \?P2\ wcode_main_tm \?Q2\" proof(rule_tac Hoare_haltI, auto) fix l r assume "wprepare_stop m args (l, r)" thus "\n. is_final (steps0 (Suc 0, l, r) wcode_main_tm n) \ (\(l, r). l = Bk # Oc # Oc \ m \ (\ln rn. r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn)) holds_for steps0 (Suc 0, l, r) wcode_main_tm n" proof(auto simp: wprepare_stop.simps) fix rn show " \n. is_final (steps0 (Suc 0, Bk # @ Bk # Bk # Oc # Oc \ m, Bk # Oc # Bk \ rn) wcode_main_tm n) \ (\(l, r). l = Bk # Oc # Oc \ m \ (\ln rn. r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn)) holds_for steps0 (Suc 0, Bk # @ Bk # Bk # Oc # Oc \ m, Bk # Oc # Bk \ rn) wcode_main_tm n" using wcode_main_tm_lemma_pre[of "args" "" 0 "Oc\(Suc m)" 0 rn,OF h refl] apply(auto simp: tape_of_nl_rev) apply(rename_tac stp ln rna) apply(rule_tac x = stp in exI, auto) done qed qed next show "composable_tm0 wcode_prepare_tm" by auto qed then obtain n where "\ tp. (case tp of (l, r) \ l = [] \ r = ) \ (is_final (steps0 (1, tp) (wcode_prepare_tm |+| wcode_main_tm) n) \ (\(l, r). \ln rn. l = Bk # Oc \ Suc m \ r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn) holds_for steps0 (1, tp) (wcode_prepare_tm |+| wcode_main_tm) n)" unfolding Hoare_halt_def by auto thus "?thesis" apply(rule_tac x = n in exI) apply(case_tac "(steps0 (Suc 0, [], ) (adjust0 wcode_prepare_tm @ shift wcode_main_tm (length wcode_prepare_tm div 2)) n)") apply(auto simp: seq_tm.simps) done qed definition tinres :: "cell list \ cell list \ bool" where "tinres xs ys = (\n. xs = ys @ Bk \ n \ ys = xs @ Bk \ n)" lemma tinres_fetch_congr[simp]: "tinres r r' \ fetch t ss (read r) = fetch t ss (read r')" apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def) using hd_replicate apply fastforce using hd_replicate apply fastforce done lemma nonempty_hd_tinres[simp]: "\tinres r r'; r \ []; r' \ []\ \ hd r = hd r'" apply(auto simp: tinres_def) done lemma tinres_nonempty[simp]: "\tinres r []; r \ []\ \ hd r = Bk" "\tinres [] r'; r' \ []\ \ hd r' = Bk" "\tinres r []; r \ []\ \ tinres (tl r) []" "tinres r r' \ tinres (b # r) (b # r')" by(auto simp: tinres_def) lemma ex_move_tl[intro]: "\na. tl r = tl (r @ Bk\(n)) @ Bk\(na) \ tl (r @ Bk\(n)) = tl r @ Bk\(na)" apply(case_tac r, simp) by(case_tac n, auto) lemma tinres_tails[simp]: "tinres r r' \ tinres (tl r) (tl r')" apply(auto simp: tinres_def) by(case_tac r', auto) lemma tinres_empty[simp]: "\tinres [] r'\ \ tinres [] (tl r')" "tinres r [] \ tinres (Bk # tl r) [Bk]" "tinres r [] \ tinres (Oc # tl r) [Oc]" by(auto simp: tinres_def) lemma tinres_step2: assumes "tinres r r'" "step0 (ss, l, r) t = (sa, la, ra)" "step0 (ss, l, r') t = (sb, lb, rb)" shows "la = lb \ tinres ra rb \ sa = sb" proof (cases "fetch t ss (read r')") case (Pair a b) have sa:"sa = sb" using assms Pair by(force simp: step.simps) have "la = lb \ tinres ra rb" using assms Pair by(cases a, auto simp: step.simps split: if_splits) thus ?thesis using sa by auto qed lemma tinres_steps2: "\tinres r r'; steps0 (ss, l, r) t stp = (sa, la, ra); steps0 (ss, l, r') t stp = (sb, lb, rb)\ \ la = lb \ tinres ra rb \ sa = sb" proof(induct stp arbitrary: sa la ra sb lb rb) case (Suc stp sa la ra sb lb rb) then show ?case apply(simp) apply(case_tac "(steps0 (ss, l, r) t stp)") apply(case_tac "(steps0 (ss, l, r') t stp)") proof - fix stp a b c aa ba ca assume ind: "\sa la ra sb lb rb. \steps0 (ss, l, r) t stp = (sa, la, ra); steps0 (ss, l, r') t stp = (sb, lb, rb)\ \ la = lb \ tinres ra rb \ sa = sb" and h: " tinres r r'" "step0 (steps0 (ss, l, r) t stp) t = (sa, la, ra)" "step0 (steps0 (ss, l, r') t stp) t = (sb, lb, rb)" "steps0 (ss, l, r) t stp = (a, b, c)" "steps0 (ss, l, r') t stp = (aa, ba, ca)" have "b = ba \ tinres c ca \ a = aa" apply(rule_tac ind, simp_all add: h) done thus "la = lb \ tinres ra rb \ sa = sb" apply(rule_tac l = b and r = c and ss = a and r' = ca and t = t in tinres_step2) using h apply(simp, simp, simp) done qed qed (simp add: steps.simps) definition wcode_adjust_tm :: "instr list" where "wcode_adjust_tm = [(WO, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), (L, 8), (L, 5), (L, 6), (WB, 5), (L, 6), (R, 7), (WO, 2), (Nop, 7), (L, 9), (WB, 8), (L, 9), (L, 10), (L, 11), (L, 10), (R, 0), (L, 11)]" lemma fetch_wcode_adjust_tm[simp]: "fetch wcode_adjust_tm (Suc 0) Bk = (WO, 1)" "fetch wcode_adjust_tm (Suc 0) Oc = (R, 2)" "fetch wcode_adjust_tm (Suc (Suc 0)) Oc = (R, 3)" "fetch wcode_adjust_tm (Suc (Suc (Suc 0))) Oc = (R, 4)" "fetch wcode_adjust_tm (Suc (Suc (Suc 0))) Bk = (R, 3)" "fetch wcode_adjust_tm 4 Bk = (L, 8)" "fetch wcode_adjust_tm 4 Oc = (L, 5)" "fetch wcode_adjust_tm 5 Oc = (WB, 5)" "fetch wcode_adjust_tm 5 Bk = (L, 6)" "fetch wcode_adjust_tm 6 Oc = (R, 7)" "fetch wcode_adjust_tm 6 Bk = (L, 6)" "fetch wcode_adjust_tm 7 Bk = (WO, 2)" "fetch wcode_adjust_tm 8 Bk = (L, 9)" "fetch wcode_adjust_tm 8 Oc = (WB, 8)" "fetch wcode_adjust_tm 9 Oc = (L, 10)" "fetch wcode_adjust_tm 9 Bk = (L, 9)" "fetch wcode_adjust_tm 10 Bk = (L, 11)" "fetch wcode_adjust_tm 10 Oc = (L, 10)" "fetch wcode_adjust_tm 11 Oc = (L, 11)" "fetch wcode_adjust_tm 11 Bk = (R, 0)" by(auto simp: fetch.simps wcode_adjust_tm_def nth_of.simps numeral_eqs_upto_12) fun wadjust_start :: "nat \ nat \ tape \ bool" where "wadjust_start m rs (l, r) = (\ ln rn. l = Bk # Oc\(Suc m) \ tl r = Oc # Bk\(ln) @ Bk # Oc\(Suc rs) @ Bk\(rn))" fun wadjust_loop_start :: "nat \ nat \ tape \ bool" where "wadjust_loop_start m rs (l, r) = (\ ln rn ml mr. l = Oc\(ml) @ Bk # Oc\(Suc m) \ r = Oc # Bk\(ln) @ Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0)" fun wadjust_loop_right_move :: "nat \ nat \ tape \ bool" where "wadjust_loop_right_move m rs (l, r) = (\ ml mr nl nr rn. l = Bk\(nl) @ Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Bk\(nr) @ Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0 \ nl + nr > 0)" fun wadjust_loop_check :: "nat \ nat \ tape \ bool" where "wadjust_loop_check m rs (l, r) = (\ ml mr ln rn. l = Oc # Bk\(ln) @ Bk # Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = (Suc rs))" fun wadjust_loop_erase :: "nat \ nat \ tape \ bool" where "wadjust_loop_erase m rs (l, r) = (\ ml mr ln rn. l = Bk\(ln) @ Bk # Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ tl r = Oc\(mr) @ Bk\(rn) \ ml + mr = (Suc rs) \ mr > 0)" fun wadjust_loop_on_left_moving_O :: "nat \ nat \ tape \ bool" where "wadjust_loop_on_left_moving_O m rs (l, r) = (\ ml mr ln rn. l = Oc\(ml) @ Bk # Oc\(Suc m )\ r = Oc # Bk\(ln) @ Bk # Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs \ mr > 0)" fun wadjust_loop_on_left_moving_B :: "nat \ nat \ tape \ bool" where "wadjust_loop_on_left_moving_B m rs (l, r) = (\ ml mr nl nr rn. l = Bk\(nl) @ Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Bk\(nr) @ Bk # Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs \ mr > 0)" fun wadjust_loop_on_left_moving :: "nat \ nat \ tape \ bool" where "wadjust_loop_on_left_moving m rs (l, r) = (wadjust_loop_on_left_moving_O m rs (l, r) \ wadjust_loop_on_left_moving_B m rs (l, r))" fun wadjust_loop_right_move2 :: "nat \ nat \ tape \ bool" where "wadjust_loop_right_move2 m rs (l, r) = (\ ml mr ln rn. l = Oc # Oc\(ml) @ Bk # Oc\(Suc m) \ r = Bk\(ln) @ Bk # Bk # Oc\(mr) @ Bk\(rn) \ ml + mr = Suc rs \ mr > 0)" fun wadjust_erase2 :: "nat \ nat \ tape \ bool" where "wadjust_erase2 m rs (l, r) = (\ ln rn. l = Bk\(ln) @ Bk # Oc # Oc\(Suc rs) @ Bk # Oc\(Suc m) \ tl r = Bk\(rn))" fun wadjust_on_left_moving_O :: "nat \ nat \ tape \ bool" where "wadjust_on_left_moving_O m rs (l, r) = (\ rn. l = Oc\(Suc rs) @ Bk # Oc\(Suc m) \ r = Oc # Bk\(rn))" fun wadjust_on_left_moving_B :: "nat \ nat \ tape \ bool" where "wadjust_on_left_moving_B m rs (l, r) = (\ ln rn. l = Bk\(ln) @ Oc # Oc\(Suc rs) @ Bk # Oc\(Suc m) \ r = Bk\(rn))" fun wadjust_on_left_moving :: "nat \ nat \ tape \ bool" where "wadjust_on_left_moving m rs (l, r) = (wadjust_on_left_moving_O m rs (l, r) \ wadjust_on_left_moving_B m rs (l, r))" fun wadjust_goon_left_moving_B :: "nat \ nat \ tape \ bool" where "wadjust_goon_left_moving_B m rs (l, r) = (\ rn. l = Oc\(Suc m) \ r = Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" fun wadjust_goon_left_moving_O :: "nat \ nat \ tape \ bool" where "wadjust_goon_left_moving_O m rs (l, r) = (\ ml mr rn. l = Oc\(ml) @ Bk # Oc\(Suc m) \ r = Oc\(mr) @ Bk\(rn) \ ml + mr = Suc (Suc rs) \ mr > 0)" fun wadjust_goon_left_moving :: "nat \ nat \ tape \ bool" where "wadjust_goon_left_moving m rs (l, r) = (wadjust_goon_left_moving_B m rs (l, r) \ wadjust_goon_left_moving_O m rs (l, r))" fun wadjust_backto_standard_pos_B :: "nat \ nat \ tape \ bool" where "wadjust_backto_standard_pos_B m rs (l, r) = (\ rn. l = [] \ r = Bk # Oc\(Suc m )@ Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" fun wadjust_backto_standard_pos_O :: "nat \ nat \ tape \ bool" where "wadjust_backto_standard_pos_O m rs (l, r) = (\ ml mr rn. l = Oc\(ml) \ r = Oc\(mr) @ Bk # Oc\(Suc (Suc rs)) @ Bk\(rn) \ ml + mr = Suc m \ mr > 0)" fun wadjust_backto_standard_pos :: "nat \ nat \ tape \ bool" where "wadjust_backto_standard_pos m rs (l, r) = (wadjust_backto_standard_pos_B m rs (l, r) \ wadjust_backto_standard_pos_O m rs (l, r))" fun wadjust_stop :: "nat \ nat \ tape \ bool" where "wadjust_stop m rs (l, r) = (\ rn. l = [Bk] \ r = Oc\(Suc m )@ Bk # Oc\(Suc (Suc rs)) @ Bk\(rn))" declare wadjust_start.simps[simp del] wadjust_loop_start.simps[simp del] wadjust_loop_right_move.simps[simp del] wadjust_loop_check.simps[simp del] wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del] wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del] wadjust_on_left_moving_O.simps[simp del] wadjust_on_left_moving_B.simps[simp del] wadjust_on_left_moving.simps[simp del] wadjust_goon_left_moving_B.simps[simp del] wadjust_goon_left_moving_O.simps[simp del] wadjust_goon_left_moving.simps[simp del] wadjust_backto_standard_pos.simps[simp del] wadjust_backto_standard_pos_B.simps[simp del] wadjust_backto_standard_pos_O.simps[simp del] wadjust_stop.simps[simp del] fun wadjust_inv :: "nat \ nat \ nat \ tape \ bool" where "wadjust_inv st m rs (l, r) = (if st = Suc 0 then wadjust_start m rs (l, r) else if st = Suc (Suc 0) then wadjust_loop_start m rs (l, r) else if st = Suc (Suc (Suc 0)) then wadjust_loop_right_move m rs (l, r) else if st = 4 then wadjust_loop_check m rs (l, r) else if st = 5 then wadjust_loop_erase m rs (l, r) else if st = 6 then wadjust_loop_on_left_moving m rs (l, r) else if st = 7 then wadjust_loop_right_move2 m rs (l, r) else if st = 8 then wadjust_erase2 m rs (l, r) else if st = 9 then wadjust_on_left_moving m rs (l, r) else if st = 10 then wadjust_goon_left_moving m rs (l, r) else if st = 11 then wadjust_backto_standard_pos m rs (l, r) else if st = 0 then wadjust_stop m rs (l, r) else False )" declare wadjust_inv.simps[simp del] fun wadjust_phase :: "nat \ config \ nat" where "wadjust_phase rs (st, l, r) = (if st = 1 then 3 else if st \ 2 \ st \ 7 then 2 else if st \ 8 \ st \ 11 then 1 else 0)" fun wadjust_stage :: "nat \ config \ nat" where "wadjust_stage rs (st, l, r) = (if st \ 2 \ st \ 7 then rs - length (takeWhile (\ a. a = Oc) (tl (dropWhile (\ a. a = Oc) (rev l @ r)))) else 0)" fun wadjust_state :: "nat \ config \ nat" where "wadjust_state rs (st, l, r) = (if st \ 2 \ st \ 7 then 8 - st else if st \ 8 \ st \ 11 then 12 - st else 0)" fun wadjust_step :: "nat \ config \ nat" where "wadjust_step rs (st, l, r) = (if st = 1 then (if hd r = Bk then 1 else 0) else if st = 3 then length r else if st = 5 then (if hd r = Oc then 1 else 0) else if st = 6 then length l else if st = 8 then (if hd r = Oc then 1 else 0) else if st = 9 then length l else if st = 10 then length l else if st = 11 then (if hd r = Bk then 0 else Suc (length l)) else 0)" fun wadjust_measure :: "(nat \ config) \ nat \ nat \ nat \ nat" where "wadjust_measure (rs, (st, l, r)) = (wadjust_phase rs (st, l, r), wadjust_stage rs (st, l, r), wadjust_state rs (st, l, r), wadjust_step rs (st, l, r))" definition wadjust_le :: "((nat \ config) \ nat \ config) set" where "wadjust_le \ (inv_image lex_square wadjust_measure)" lemma wf_lex_square[intro]: "wf lex_square" by(auto simp: Abacus.lex_pair_def lex_square_def Abacus.lex_triple_def) lemma wf_wadjust_le[intro]: "wf wadjust_le" by(auto simp: wadjust_le_def Abacus.lex_triple_def Abacus.lex_pair_def) lemma wadjust_start_snd_nonempty[simp]: "wadjust_start m rs (c, []) = False" apply(auto simp: wadjust_start.simps) done lemma wadjust_loop_right_move_fst_nonempty[simp]: "wadjust_loop_right_move m rs (c, []) \ c \ []" apply(auto simp: wadjust_loop_right_move.simps) done lemma wadjust_loop_check_fst_nonempty[simp]: "wadjust_loop_check m rs (c, []) \ c \ []" apply(simp only: wadjust_loop_check.simps, auto) done lemma wadjust_loop_start_snd_nonempty[simp]: "wadjust_loop_start m rs (c, []) = False" apply(simp add: wadjust_loop_start.simps) done lemma wadjust_erase2_singleton[simp]: "wadjust_loop_check m rs (c, []) \ wadjust_erase2 m rs (tl c, [hd c])" apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto) done lemma wadjust_loop_on_left_moving_snd_nonempty[simp]: "wadjust_loop_on_left_moving m rs (c, []) = False" "wadjust_loop_right_move2 m rs (c, []) = False" "wadjust_erase2 m rs ([], []) = False" by(auto simp: wadjust_loop_on_left_moving.simps wadjust_loop_right_move2.simps wadjust_erase2.simps) lemma wadjust_on_left_moving_B_Bk1[simp]: "wadjust_on_left_moving_B m rs (Oc # Oc # Oc\(rs) @ Bk # Oc # Oc\(m), [Bk])" apply(simp add: wadjust_on_left_moving_B.simps, auto) done lemma wadjust_on_left_moving_B_Bk2[simp]: "wadjust_on_left_moving_B m rs (Bk\(n) @ Bk # Oc # Oc # Oc\(rs) @ Bk # Oc # Oc\(m), [Bk])" apply(simp add: wadjust_on_left_moving_B.simps , auto) apply(rule_tac x = "Suc n" in exI, simp add: exp_ind del: replicate_Suc) done lemma wadjust_on_left_moving_singleton[simp]: "\wadjust_erase2 m rs (c, []); c \ []\ \ wadjust_on_left_moving m rs (tl c, [hd c])" unfolding wadjust_erase2.simps apply(auto simp add: wadjust_on_left_moving.simps) apply (metis (no_types, lifting) empty_replicate hd_append hd_replicate list.sel(1) list.sel(3) self_append_conv2 tl_append2 tl_replicate wadjust_on_left_moving_B_Bk1 wadjust_on_left_moving_B_Bk2)+ done lemma wadjust_erase2_cases[simp]: "wadjust_erase2 m rs (c, []) \ (c = [] \ wadjust_on_left_moving m rs ([], [Bk])) \ (c \ [] \ wadjust_on_left_moving m rs (tl c, [hd c]))" apply(auto) done lemma wadjust_on_left_moving_nonempty[simp]: "wadjust_on_left_moving m rs ([], []) = False" "wadjust_on_left_moving_O m rs (c, []) = False" apply(auto simp: wadjust_on_left_moving.simps wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) done lemma wadjust_on_left_moving_B_singleton_Bk[simp]: " \wadjust_on_left_moving_B m rs (c, []); c \ []; hd c = Bk\ \ wadjust_on_left_moving_B m rs (tl c, [Bk])" apply(auto simp add: wadjust_on_left_moving_B.simps hd_append) by (metis cell.distinct(1) empty_replicate list.sel(1) tl_append2 tl_replicate) lemma wadjust_on_left_moving_B_singleton_Oc[simp]: "\wadjust_on_left_moving_B m rs (c, []); c \ []; hd c = Oc\ \ wadjust_on_left_moving_O m rs (tl c, [Oc])" apply(auto simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps hd_append) apply (metis cell.distinct(1) empty_replicate hd_replicate list.sel(3) self_append_conv2)+ done lemma wadjust_on_left_moving_singleton2[simp]: "\wadjust_on_left_moving m rs (c, []); c \ []\ \ wadjust_on_left_moving m rs (tl c, [hd c])" apply(simp add: wadjust_on_left_moving.simps) apply(case_tac "hd c", simp_all) done lemma wadjust_nonempty[simp]: "wadjust_goon_left_moving m rs (c, []) = False" "wadjust_backto_standard_pos m rs (c, []) = False" by(auto simp: wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps wadjust_goon_left_moving_O.simps wadjust_backto_standard_pos.simps wadjust_backto_standard_pos_B.simps wadjust_backto_standard_pos_O.simps) lemma wadjust_loop_start_no_Bk[simp]: "wadjust_loop_start m rs (c, Bk # list) = False" apply(auto simp: wadjust_loop_start.simps) done lemma wadjust_loop_check_nonempty[simp]: "wadjust_loop_check m rs (c, b) \ c \ []" apply(simp only: wadjust_loop_check.simps, auto) done lemma wadjust_erase2_via_loop_check_Bk[simp]: "wadjust_loop_check m rs (c, Bk # list) \ wadjust_erase2 m rs (tl c, hd c # Bk # list)" by (auto simp: wadjust_loop_check.simps wadjust_erase2.simps) declare wadjust_loop_on_left_moving_O.simps[simp del] wadjust_loop_on_left_moving_B.simps[simp del] lemma wadjust_loop_on_left_moving_B_via_erase[simp]: "\wadjust_loop_erase m rs (c, Bk # list); hd c = Bk\ \ wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)" unfolding wadjust_loop_erase.simps wadjust_loop_on_left_moving_B.simps apply(erule_tac exE)+ apply(rename_tac ml mr ln rn) apply(rule_tac x = ml in exI, rule_tac x = mr in exI, rule_tac x = ln in exI, rule_tac x = 0 in exI) apply(case_tac ln, auto) apply(simp add: exp_ind [THEN sym]) done lemma wadjust_loop_on_left_moving_O_Bk_via_erase[simp]: "\wadjust_loop_erase m rs (c, Bk # list); c \ []; hd c = Oc\ \ wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)" apply(auto simp: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps) by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(1)) lemma wadjust_loop_on_left_moving_Bk_via_erase[simp]: "\wadjust_loop_erase m rs (c, Bk # list); c \ []\ \ wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)" apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps) done lemma wadjust_loop_on_left_moving_B_Bk_move[simp]: "\wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\ \ wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)" apply(simp only: wadjust_loop_on_left_moving_B.simps) apply(erule_tac exE)+ by (metis (no_types, lifting) cell.distinct(1) list.sel(1) replicate_Suc_iff_anywhere self_append_conv2 tl_append2 tl_replicate) lemma wadjust_loop_on_left_moving_O_Oc_move[simp]: "\wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\ \ wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)" apply(simp only: wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps) by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(3) self_append_conv2) lemma wadjust_loop_erase_nonempty[simp]: "wadjust_loop_erase m rs (c, b) \ c \ []" "wadjust_loop_on_left_moving m rs (c, b) \ c \ []" "wadjust_loop_right_move2 m rs (c, b) \ c \ []" "wadjust_erase2 m rs (c, Bk # list) \ c \ []" "wadjust_on_left_moving m rs (c,b) \ c \ []" "wadjust_on_left_moving_O m rs (c, Bk # list) = False" "wadjust_goon_left_moving m rs (c, b) \ c \ []" "wadjust_loop_on_left_moving_O m rs (c, Bk # list) = False" by(auto simp: wadjust_loop_erase.simps wadjust_loop_on_left_moving.simps wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps wadjust_loop_right_move2.simps wadjust_erase2.simps wadjust_on_left_moving.simps wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps wadjust_goon_left_moving_O.simps) lemma wadjust_loop_on_left_moving_Bk_move[simp]: "wadjust_loop_on_left_moving m rs (c, Bk # list) \ wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)" apply(simp add: wadjust_loop_on_left_moving.simps) apply(case_tac "hd c", simp_all) done lemma wadjust_loop_start_Oc_via_Bk_move[simp]: "wadjust_loop_right_move2 m rs (c, Bk # list) \ wadjust_loop_start m rs (c, Oc # list)" apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps replicate_app_Cons_same) by (metis add_Suc replicate_Suc) lemma wadjust_on_left_moving_Bk_via_erase[simp]: "wadjust_erase2 m rs (c, Bk # list) \ wadjust_on_left_moving m rs (tl c, hd c # Bk # list)" apply(auto simp: wadjust_erase2.simps wadjust_on_left_moving.simps replicate_app_Cons_same wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) apply (metis exp_ind replicate_append_same)+ done lemma wadjust_on_left_moving_B_Bk_drop_one: "\wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\ \ wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)" apply(auto simp: wadjust_on_left_moving_B.simps) by (metis cell.distinct(1) hd_append list.sel(1) tl_append2 tl_replicate) lemma wadjust_on_left_moving_B_Bk_drop_Oc: "\wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\ \ wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)" apply(auto simp: wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps) by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(3) self_append_conv2) lemma wadjust_on_left_moving_B_drop[simp]: "wadjust_on_left_moving m rs (c, Bk # list) \ wadjust_on_left_moving m rs (tl c, hd c # Bk # list)" by(cases "hd c", auto simp:wadjust_on_left_moving.simps wadjust_on_left_moving_B_Bk_drop_one wadjust_on_left_moving_B_Bk_drop_Oc) lemma wadjust_goon_left_moving_O_no_Bk[simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False" by (auto simp add: wadjust_goon_left_moving_O.simps) lemma wadjust_backto_standard_pos_via_left_Bk[simp]: "wadjust_goon_left_moving m rs (c, Bk # list) \ wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)" by(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps wadjust_backto_standard_pos_O.simps) lemma wadjust_loop_right_move_Oc[simp]: "wadjust_loop_start m rs (c, Oc # list) \ wadjust_loop_right_move m rs (Oc # c, list)" apply(auto simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps simp del:split_head_repeat) apply(rename_tac ln rn ml mr) apply(rule_tac x = ml in exI, rule_tac x = mr in exI, rule_tac x = 0 in exI, simp) apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind del: replicate_Suc) done lemma wadjust_loop_check_Oc[simp]: assumes "wadjust_loop_right_move m rs (c, Oc # list)" shows "wadjust_loop_check m rs (Oc # c, list)" proof - from assms obtain ml mr nl nr rn where "c = Bk \ nl @ Oc # Oc \ ml @ Bk # Oc \ m @ [Oc]" "Oc # list = Bk \ nr @ Oc \ mr @ Bk \ rn" "ml + mr = Suc (Suc rs)" "0 < mr" "0 < nl + nr" unfolding wadjust_loop_right_move.simps exp_ind wadjust_loop_check.simps by auto hence "\ln. Oc # c = Oc # Bk \ ln @ Bk # Oc # Oc \ ml @ Bk # Oc \ Suc m" "\rn. list = Oc \ (mr - 1) @ Bk \ rn" "ml + (mr - 1) = Suc rs" by(cases nl;cases nr;cases mr;force simp add: wadjust_loop_right_move.simps exp_ind wadjust_loop_check.simps replicate_append_same)+ thus ?thesis unfolding wadjust_loop_check.simps by auto qed lemma wadjust_loop_erase_move_Oc[simp]: "wadjust_loop_check m rs (c, Oc # list) \ wadjust_loop_erase m rs (tl c, hd c # Oc # list)" apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps) apply(erule_tac exE)+ using Cons_replicate_eq by fastforce lemma wadjust_loop_on_move_no_Oc[simp]: "wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False" "wadjust_loop_right_move2 m rs (c, Oc # list) = False" "wadjust_loop_on_left_moving m rs (c, Oc # list) \ wadjust_loop_right_move2 m rs (Oc # c, list)" "wadjust_on_left_moving_B m rs (c, Oc # list) = False" "wadjust_loop_erase m rs (c, Oc # list) \ wadjust_loop_erase m rs (c, Bk # list)" by(auto simp: wadjust_loop_on_left_moving_B.simps wadjust_loop_on_left_moving_O.simps wadjust_loop_right_move2.simps replicate_app_Cons_same wadjust_loop_on_left_moving.simps wadjust_on_left_moving_B.simps wadjust_loop_erase.simps) lemma wadjust_goon_left_moving_B_Bk_Oc: "\wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\ \ wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)" apply(auto simp: wadjust_on_left_moving_O.simps wadjust_goon_left_moving_B.simps ) done lemma wadjust_goon_left_moving_O_Oc_Oc: "\wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\ \ wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)" apply(auto simp: wadjust_on_left_moving_O.simps wadjust_goon_left_moving_O.simps ) apply(auto simp: numeral_2_eq_2) done lemma wadjust_goon_left_moving_Oc[simp]: "wadjust_on_left_moving m rs (c, Oc # list) \ wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)" by(cases "hd c"; force simp: wadjust_on_left_moving.simps wadjust_goon_left_moving.simps wadjust_goon_left_moving_B_Bk_Oc wadjust_goon_left_moving_O_Oc_Oc)+ lemma left_moving_Bk_Oc[simp]: "\wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\ \ wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)" apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps hd_append dest!: gr0_implies_Suc) apply (metis cell.distinct(1) empty_replicate hd_replicate list.sel(3) self_append_conv2) by (metis add_cancel_right_left cell.distinct(1) hd_replicate replicate_Suc_iff_anywhere) lemma left_moving_Oc_Oc[simp]: "\wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\ \ wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)" apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps) apply(rename_tac mlx mrx rnx) apply(rule_tac x = "mlx - 1" in exI, simp) apply(case_tac mlx, simp_all add: ) apply(rule_tac x = "Suc mrx" in exI, auto simp: ) done lemma wadjust_goon_left_moving_B_no_Oc[simp]: "wadjust_goon_left_moving_B m rs (c, Oc # list) = False" apply(auto simp: wadjust_goon_left_moving_B.simps) done lemma wadjust_goon_left_moving_Oc_move[simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \ wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)" by(cases "hd c",auto simp: wadjust_goon_left_moving.simps) lemma wadjust_backto_standard_pos_B_no_Oc[simp]: "wadjust_backto_standard_pos_B m rs (c, Oc # list) = False" apply(simp add: wadjust_backto_standard_pos_B.simps) done lemma wadjust_backto_standard_pos_O_no_Bk[simp]: "wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False" by(simp add: wadjust_backto_standard_pos_O.simps) lemma wadjust_backto_standard_pos_B_Bk_Oc[simp]: "wadjust_backto_standard_pos_O m rs ([], Oc # list) \ wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)" apply(auto simp: wadjust_backto_standard_pos_O.simps wadjust_backto_standard_pos_B.simps) done lemma wadjust_backto_standard_pos_B_Bk_Oc_via_O[simp]: "\wadjust_backto_standard_pos_O m rs (c, Oc # list); c \ []; hd c = Bk\ \ wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)" apply(simp add:wadjust_backto_standard_pos_O.simps wadjust_backto_standard_pos_B.simps, auto) done lemma wadjust_backto_standard_pos_B_Oc_Oc_via_O[simp]: "\wadjust_backto_standard_pos_O m rs (c, Oc # list); c \ []; hd c = Oc\ \ wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)" apply(simp add: wadjust_backto_standard_pos_O.simps, auto) by force lemma wadjust_backto_standard_pos_cases[simp]: "wadjust_backto_standard_pos m rs (c, Oc # list) \ (c = [] \ wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \ (c \ [] \ wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))" apply(auto simp: wadjust_backto_standard_pos.simps) apply(case_tac "hd c", simp_all) done lemma wadjust_loop_right_move_nonempty_snd[simp]: "wadjust_loop_right_move m rs (c, []) = False" proof - {fix nl ml mr rn nr have "(c = Bk \ nl @ Oc # Oc \ ml @ Bk # Oc \ Suc m \ [] = Bk \ nr @ Oc \ mr @ Bk \ rn \ ml + mr = Suc (Suc rs) \ 0 < mr \ 0 < nl + nr) = False" by auto } note t=this thus ?thesis unfolding wadjust_loop_right_move.simps t by blast qed lemma wadjust_loop_erase_nonempty_snd[simp]: "wadjust_loop_erase m rs (c, []) = False" apply(simp only: wadjust_loop_erase.simps, auto) done lemma wadjust_loop_erase_cases2[simp]: "\Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Bk # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" apply(simp only: wadjust_loop_erase.simps) apply(rule_tac disjI2) apply(case_tac c, simp, simp) done lemma dropWhile_exp1: "dropWhile (\a. a = Oc) (Oc\(n) @ xs) = dropWhile (\a. a = Oc) xs" apply(induct n, simp_all add: ) done lemma takeWhile_exp1: "takeWhile (\a. a = Oc) (Oc\(n) @ xs) = Oc\(n) @ takeWhile (\a. a = Oc) xs" apply(induct n, simp_all add: ) done lemma wadjust_correctness_helper_1: assumes "Suc (Suc rs) = a" " wadjust_loop_right_move2 m rs (c, Bk # list)" shows "a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" proof - have "ml + mr = Suc rs \ 0 < mr \ rs - (ml + length (takeWhile (\a. a = Oc) list)) < Suc rs - (ml + length (takeWhile (\a. a = Oc) (Bk \ ln @ Bk # Bk # Oc \ mr @ Bk \ rn))) " for ml mr ln rn by(cases ln, auto) thus ?thesis using assms by (auto simp: wadjust_loop_right_move2.simps dropWhile_exp1 takeWhile_exp1) qed lemma wadjust_correctness_helper_2: "\Suc (Suc rs) = a; wadjust_loop_on_left_moving m rs (c, Bk # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list))))" apply(subgoal_tac "c \ []") apply(case_tac c, simp_all) done lemma wadjust_loop_check_empty_false[simp]: "wadjust_loop_check m rs ([], b) = False" apply(simp add: wadjust_loop_check.simps) done lemma wadjust_loop_check_cases: "\Suc (Suc rs) = a; wadjust_loop_check m rs (c, Oc # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list))))" apply(case_tac "c", simp_all) done lemma wadjust_loop_erase_cases_or: "\Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Oc # list)\ \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) < a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list)))) \ a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Bk # list)))) = a - length (takeWhile (\a. a = Oc) (tl (dropWhile (\a. a = Oc) (rev c @ Oc # list))))" apply(simp add: wadjust_loop_erase.simps) apply(rule_tac disjI2) apply(auto) apply(simp add: dropWhile_exp1 takeWhile_exp1) done lemmas wadjust_correctness_helpers = wadjust_correctness_helper_2 wadjust_correctness_helper_1 wadjust_loop_erase_cases_or wadjust_loop_check_cases declare numeral_2_eq_2[simp del] lemma wadjust_start_Oc[simp]: "wadjust_start m rs (c, Bk # list) \ wadjust_start m rs (c, Oc # list)" apply(auto simp: wadjust_start.simps) done lemma wadjust_stop_Bk[simp]: "wadjust_backto_standard_pos m rs (c, Bk # list) \ wadjust_stop m rs (Bk # c, list)" apply(auto simp: wadjust_backto_standard_pos.simps wadjust_stop.simps wadjust_backto_standard_pos_B.simps) done lemma wadjust_loop_start_Oc[simp]: assumes "wadjust_start m rs (c, Oc # list)" shows "wadjust_loop_start m rs (Oc # c, list)" proof - from assms[unfolded wadjust_start.simps] obtain ln rn where "c = Bk # Oc # Oc \ m" "list = Oc # Bk \ ln @ Bk # Oc # Oc \ rs @ Bk \ rn" by(auto) hence "Oc # c = Oc \ 1 @ Bk # Oc \ Suc m \ list = Oc # Bk \ ln @ Bk # Oc \Suc rs @ Bk \ rn \ 1 + (Suc rs) = Suc (Suc rs) \ 0 < Suc rs" by auto thus ?thesis unfolding wadjust_loop_start.simps by blast qed lemma erase2_Bk_if_Oc[simp]:" wadjust_erase2 m rs (c, Oc # list) \ wadjust_erase2 m rs (c, Bk # list)" apply(auto simp: wadjust_erase2.simps) done lemma wadjust_loop_right_move_Bk[simp]: "wadjust_loop_right_move m rs (c, Bk # list) \ wadjust_loop_right_move m rs (Bk # c, list)" apply(simp only: wadjust_loop_right_move.simps) apply(erule_tac exE)+ apply auto apply (metis cell.distinct(1) empty_replicate hd_append hd_replicate less_SucI list.sel(1) list.sel(3) neq0_conv replicate_Suc_iff_anywhere tl_append2 tl_replicate)+ done lemma wadjust_correctness: shows "let P = (\ (len, st, l, r). st = 0) in let Q = (\ (len, st, l, r). wadjust_inv st m rs (l, r)) in let f = (\ stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\(Suc m), Bk # Oc # Bk\(ln) @ Bk # Oc\(Suc rs) @ Bk\(rn)) wcode_adjust_tm stp)) in \ n .P (f n) \ Q (f n)" proof - let ?P = "(\ (len, st, l, r). st = 0)" let ?Q = "\ (len, st, l, r). wadjust_inv st m rs (l, r)" let ?f = "\ stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\(Suc m), Bk # Oc # Bk\(ln) @ Bk # Oc\(Suc rs) @ Bk\(rn)) wcode_adjust_tm stp)" have "\ n. ?P (?f n) \ ?Q (?f n)" proof(rule_tac halt_lemma2) show "wf wadjust_le" by auto next { fix n assume a:"\ ?P (?f n) \ ?Q (?f n)" have "?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wadjust_le" proof(cases "?f n") case (fields a b c d) then show ?thesis proof(cases d) case Nil then show ?thesis using a fields apply(simp add: step.simps) apply(simp_all only: wadjust_inv.simps split: if_splits) apply(simp_all add: wadjust_inv.simps wadjust_le_def wadjust_correctness_helpers Abacus.lex_triple_def Abacus.lex_pair_def lex_square_def split: if_splits). next case (Cons aa list) then show ?thesis using a fields Nil Cons apply((case_tac aa); simp add: step.simps) apply(simp_all only: wadjust_inv.simps split: if_splits) apply(simp_all) apply(simp_all add: wadjust_inv.simps wadjust_le_def wadjust_correctness_helpers Abacus.lex_triple_def Abacus.lex_pair_def lex_square_def split: if_splits). qed qed } thus "\ n. \ ?P (?f n) \ ?Q (?f n) \ ?Q (?f (Suc n)) \ (?f (Suc n), ?f n) \ wadjust_le" by auto next show "?Q (?f 0)" by(auto simp add: steps.simps wadjust_inv.simps wadjust_start.simps) next show "\ ?P (?f 0)" by (simp add: steps.simps) qed thus"?thesis" by simp qed lemma composable_tm_wcode_adjust_tm[intro]: "composable_tm (wcode_adjust_tm, 0)" by(auto simp: wcode_adjust_tm_def composable_tm.simps) lemma bl_bin_nonzero[simp]: "args \ [] \ bl_bin () > 0" by(cases args) (auto simp: tape_of_nl_cons bl_bin.simps) lemma wcode_lemma_pre': "args \ [] \ \ stp rn. steps0 (Suc 0, [], ) ((wcode_prepare_tm |+| wcode_main_tm) |+| wcode_adjust_tm) stp = (0, [Bk], Oc\(Suc m) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn))" proof - let ?P1 = "\ (l, r). l = [] \ r = " let ?Q1 = "\(l, r). l = Bk # Oc\(Suc m) \ (\ln rn. r = Bk # Oc # Bk\(ln) @ Bk # Bk # Oc\(bl_bin ()) @ Bk\(rn))" let ?P2 = ?Q1 let ?Q2 = "\ (l, r). (wadjust_stop m (bl_bin () - 1) (l, r))" let ?P3 = "\ tp. False" assume h: "args \ []" hence a: "bl_bin () > 0" using h by simp hence "\?P1\ (wcode_prepare_tm |+| wcode_main_tm) |+| wcode_adjust_tm \?Q2\" proof(rule_tac Hoare_plus_halt) next show "composable_tm (wcode_prepare_tm |+| wcode_main_tm, 0)" by(rule_tac seq_tm_composable, auto) next show "\?P1\ wcode_prepare_tm |+| wcode_main_tm \?Q1\" proof(rule_tac Hoare_haltI, auto) show "\n. is_final (steps0 (Suc 0, [], ) (wcode_prepare_tm |+| wcode_main_tm) n) \ (\(l, r). l = Bk # Oc # Oc \ m \ (\ln rn. r = Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn)) holds_for steps0 (Suc 0, [], ) (wcode_prepare_tm |+| wcode_main_tm) n" using h prepare_mainpart_lemma[of args m] apply(auto) apply(rename_tac stp ln rn) apply(rule_tac x = stp in exI, simp) apply(rule_tac x = ln in exI, auto) done qed next show "\?P2\ wcode_adjust_tm \?Q2\" proof(rule_tac Hoare_haltI, auto del: replicate_Suc) fix ln rn obtain n a b where "steps0 (Suc 0, Bk # Oc \ m @ [Oc], Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ (bl_bin () - Suc 0) @ Oc # Bk \ rn) wcode_adjust_tm n = (0, a, b)" "wadjust_inv 0 m (bl_bin () - Suc 0) (a, b)" using wadjust_correctness[of m "bl_bin () - 1" "Suc ln" rn,unfolded Let_def] by(simp del: replicate_Suc add: replicate_Suc[THEN sym] exp_ind, auto) thus "\n. is_final (steps0 (Suc 0, Bk # Oc # Oc \ m, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn) wcode_adjust_tm n) \ wadjust_stop m (bl_bin () - Suc 0) holds_for steps0 (Suc 0, Bk # Oc # Oc \ m, Bk # Oc # Bk \ ln @ Bk # Bk # Oc \ bl_bin () @ Bk \ rn) wcode_adjust_tm n" apply(rule_tac x = n in exI) using a apply(case_tac "bl_bin ()", simp, simp del: replicate_Suc add: exp_ind wadjust_inv.simps) by (simp add: replicate_append_same) qed qed thus "?thesis" apply(simp add: Hoare_halt_def, auto) apply(rename_tac n) apply(case_tac "(steps0 (Suc 0, [], <(m::nat) # args>) ((wcode_prepare_tm |+| wcode_main_tm) |+| wcode_adjust_tm) n)") apply(rule_tac x = n in exI, auto simp: wadjust_stop.simps) using a apply(case_tac "bl_bin ()", simp_all) done qed text \ The initialization TM \wcode_tm\. \ definition wcode_tm :: "instr list" where "wcode_tm = (wcode_prepare_tm |+| wcode_main_tm) |+| wcode_adjust_tm" text \ The correctness of \wcode_tm\. \ lemma wcode_lemma_1: "args \ [] \ \ stp ln rn. steps0 (Suc 0, [], ) (wcode_tm) stp = (0, [Bk], Oc\(Suc m) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn))" apply(simp add: wcode_lemma_pre' wcode_tm_def del: replicate_Suc) done lemma wcode_lemma: "args \ [] \ \ stp ln rn. steps0 (Suc 0, [], ) (wcode_tm) stp = (0, [Bk], <[m ,bl_bin ()]> @ Bk\(rn))" using wcode_lemma_1[of args m] apply(simp add: wcode_tm_def tape_of_list_def tape_of_nat_def) done section \The Universal TM\ text \ This section gives the explicit construction of {\em Universal Turing Machine}, defined as \utm\ and proves its correctness. It is pretty easy by composing the partial results we have got so far. \ subsection \Definition of the machine utm\ definition utm :: "instr list" where "utm = (let (aprog, rs_pos, a_md) = rec_ci rec_F in let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in (wcode_tm |+| (tm_of abc_F @ shift (mopup_n_tm (Suc (Suc 0))) (length (tm_of abc_F) div 2))))" definition f_aprog :: "abc_prog" where "f_aprog \ (let (aprog, rs_pos, a_md) = rec_ci rec_F in aprog [+] dummy_abc (Suc (Suc 0)))" definition f_tprog_tm :: "instr list" where "f_tprog_tm = tm_of (f_aprog)" definition utm_with_two_args :: "instr list" where "utm_with_two_args \ f_tprog_tm @ shift (mopup_n_tm (Suc (Suc 0))) (length f_tprog_tm div 2)" definition utm_pre_tm :: "instr list" where "utm_pre_tm = wcode_tm |+| utm_with_two_args" (* SPIKE FABR: some lemmas for clarification *) lemma fabr_spike_1: "utm_with_two_args = tm_of (fst (rec_ci rec_F) [+] dummy_abc (Suc (Suc 0))) @ shift (mopup_n_tm (Suc (Suc 0))) (length (tm_of (fst (rec_ci rec_F) [+] dummy_abc (Suc (Suc 0)))) div 2)" proof (cases "rec_ci rec_F") case (fields a b c) then show ?thesis by (simp add: fields f_aprog_def f_tprog_tm_def utm_with_two_args_def) qed lemma fabr_spike_2: "utm = wcode_tm |+| tm_of (fst (rec_ci rec_F) [+] dummy_abc (Suc (Suc 0))) @ shift (mopup_n_tm (Suc (Suc 0))) (length (tm_of (fst (rec_ci rec_F) [+] dummy_abc (Suc (Suc 0)))) div 2)" proof (cases "rec_ci rec_F") case (fields a b c) then show ?thesis by (simp add: fields f_aprog_def f_tprog_tm_def utm_def) qed theorem fabr_spike_3: "utm = wcode_tm |+| utm_with_two_args" using fabr_spike_1 fabr_spike_2 by auto corollary fabr_spike_4: "utm = utm_pre_tm" using fabr_spike_3 utm_pre_tm_def by auto (* END SPIKE *) lemma tinres_step1: assumes "tinres l l'" "step (ss, l, r) (t, 0) = (sa, la, ra)" "step (ss, l', r) (t, 0) = (sb, lb, rb)" shows "tinres la lb \ ra = rb \ sa = sb" proof(cases "r") case Nil then show ?thesis using assms by (cases "(fetch t ss Bk)";cases "fst (fetch t ss Bk)";auto simp:step.simps split:if_splits) next case (Cons a list) then show ?thesis using assms by (cases "(fetch t ss a)";cases "fst (fetch t ss a)";auto simp:step.simps split:if_splits) qed lemma tinres_steps1: "\tinres l l'; steps (ss, l, r) (t, 0) stp = (sa, la, ra); steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\ \ tinres la lb \ ra = rb \ sa = sb" proof (induct stp arbitrary: sa la ra sb lb rb) case (Suc stp) then show ?case apply simp apply(case_tac "(steps (ss, l, r) (t, 0) stp)") apply(case_tac "(steps (ss, l', r) (t, 0) stp)") proof - fix stp sa la ra sb lb rb a b c aa ba ca assume ind: "\sa la ra sb lb rb. \steps (ss, l, r) (t, 0) stp = (sa, (la::cell list), ra); steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\ \ tinres la lb \ ra = rb \ sa = sb" and h: " tinres l l'" "step (steps (ss, l, r) (t, 0) stp) (t, 0) = (sa, la, ra)" "step (steps (ss, l', r) (t, 0) stp) (t, 0) = (sb, lb, rb)" "steps (ss, l, r) (t, 0) stp = (a, b, c)" "steps (ss, l', r) (t, 0) stp = (aa, ba, ca)" have "tinres b ba \ c = ca \ a = aa" using ind h by metis thus "tinres la lb \ ra = rb \ sa = sb" using tinres_step1 h by metis qed qed (simp add: steps.simps) lemma tinres_some_exp[simp]: "tinres (Bk \ m @ [Bk, Bk]) la \ \m. la = Bk \ m" unfolding tinres_def proof - let ?c1 = "\ n. Bk \ m @ [Bk, Bk] = la @ Bk \ n" let ?c2 = "\ n. la = (Bk \ m @ [Bk, Bk]) @ Bk \ n" assume "\n. ?c1 n \ ?c2 n" then obtain n where "?c1 n \ ?c2 n" by auto then consider "?c1 n" | "?c2 n" by blast thus ?thesis proof(cases) case 1 hence "Bk \ Suc (Suc m) = la @ Bk \ n" by (metis exp_ind append_Cons append_eq_append_conv2 self_append_conv2) hence "la = Bk \ (Suc (Suc m) - n)" by (metis replicate_add append_eq_append_conv diff_add_inverse2 length_append length_replicate) then show ?thesis by auto next case 2 hence "la = Bk \ (m + Suc (Suc n))" by (metis append_Cons append_eq_append_conv2 replicate_Suc replicate_add self_append_conv2) then show ?thesis by blast qed qed lemma utm_with_two_args_halt_eq: assumes composable_tm: "composable_tm (tp, 0)" and exec: "steps0 (Suc 0, Bk\(l), ) tp stp = (0, Bk\(m), Oc\(rs)@Bk\(n))" and resutl: "0 < rs" shows "\stp m n. steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\(i)) utm_with_two_args stp = (0, Bk\(m), Oc\(rs) @ Bk\(n))" proof - obtain ap arity fp where a: "rec_ci rec_F = (ap, arity, fp)" by (metis prod_cases3) moreover have b: "rec_exec rec_F [code tp, (bl2wc ())] = (rs - Suc 0)" using assms apply(rule_tac F_correct, simp_all) done have "\ stp m l. steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc ()]> @ Bk\i) (f_tprog_tm @ shift (mopup_n_tm (length [code tp, bl2wc ()])) (length f_tprog_tm div 2)) stp = (0, Bk\m @ Bk # Bk # [], Oc\Suc (rec_exec rec_F [code tp, (bl2wc ())]) @ Bk\l)" proof(rule_tac recursive_compile_to_tm_correct1) show "rec_ci rec_F = (ap, arity, fp)" using a by simp next show "terminate rec_F [code tp, bl2wc ()]" using assms by(rule_tac terminate_F, simp_all) next show "f_tprog_tm = tm_of (ap [+] dummy_abc (length [code tp, bl2wc ()]))" using a apply(simp add: f_tprog_tm_def f_aprog_def numeral_2_eq_2) done qed then obtain stp m l where "steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc ()]> @ Bk\i) (f_tprog_tm @ shift (mopup_n_tm (length [code tp, (bl2wc ())])) (length f_tprog_tm div 2)) stp = (0, Bk\m @ Bk # Bk # [], Oc\Suc (rec_exec rec_F [code tp, (bl2wc ())]) @ Bk\l)" by blast hence "\ m. steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\i) (f_tprog_tm @ shift (mopup_n_tm 2) (length f_tprog_tm div 2)) stp = (0, Bk\m, Oc\Suc (rs - 1) @ Bk\l)" proof - assume g: "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]> @ Bk \ i) (f_tprog_tm @ shift (mopup_n_tm (length [code tp, bl2wc ()])) (length f_tprog_tm div 2)) stp = (0, Bk \ m @ [Bk, Bk], Oc \ Suc ((rec_exec rec_F [code tp, bl2wc ()])) @ Bk \ l)" moreover have "tinres [Bk, Bk] [Bk]" apply(auto simp: tinres_def) done moreover obtain sa la ra where "steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\i) (f_tprog_tm @ shift (mopup_n_tm 2) (length f_tprog_tm div 2)) stp = (sa, la, ra)" apply(case_tac "steps0 (Suc 0, [Bk], <[code tp, bl2wc ()]> @ Bk\i) (f_tprog_tm @ shift (mopup_n_tm 2) (length f_tprog_tm div 2)) stp", auto) done ultimately show "?thesis" using b apply(drule_tac la = "Bk\m @ [Bk, Bk]" in tinres_steps1, auto simp: numeral_2_eq_2) done qed thus "?thesis" apply(auto) apply(rule_tac x = stp in exI, simp add: utm_with_two_args_def) using assms apply(case_tac rs, simp_all add: numeral_2_eq_2) done qed lemma composable_tm_wcode_tm[intro]: "composable_tm (wcode_tm, 0)" apply(simp add: wcode_tm_def) apply(rule_tac seq_tm_composable) apply(rule_tac seq_tm_composable, auto) done lemma utm_halt_lemma_pre: assumes "composable_tm (tp, 0)" and result: "0 < rs" and args: "args \ []" and exec: "steps0 (Suc 0, Bk\(i), ) tp stp = (0, Bk\(m), Oc\(rs)@Bk\(k))" shows "\stp m n. steps0 (Suc 0, [], ) utm_pre_tm stp = (0, Bk\(m), Oc\(rs) @ Bk\(n))" proof - let ?Q2 = "\ (l, r). (\ ln rn. l = Bk\(ln) \ r = Oc\(rs) @ Bk\(rn))" let ?P1 = "\ (l, r). l = [] \ r = " let ?Q1 = "\ (l, r). (l = [Bk] \ (\ rn. r = Oc\(Suc (code tp)) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn)))" let ?P2 = ?Q1 let ?P3 = "\ (l, r). False" have "\?P1\ (wcode_tm |+| utm_with_two_args) \?Q2\" proof(rule_tac Hoare_plus_halt) show "composable_tm (wcode_tm, 0)" by auto next show "\?P1\ wcode_tm \?Q1\" apply(rule_tac Hoare_haltI, auto) using wcode_lemma_1[of args "code tp"] args apply(auto) by (metis (mono_tags, lifting) holds_for.simps is_finalI old.prod.case) next show "\?P2\ utm_with_two_args \?Q2\" proof(rule_tac Hoare_haltI, auto) fix rn show "\n. is_final (steps0 (Suc 0, [Bk], Oc # Oc \ code tp @ Bk # Oc # Oc \ bl_bin () @ Bk \ rn) utm_with_two_args n) \ (\(l, r). (\ln. l = Bk \ ln) \ (\rn. r = Oc \ rs @ Bk \ rn)) holds_for steps0 (Suc 0, [Bk], Oc # Oc \ code tp @ Bk # Oc # Oc \ bl_bin () @ Bk \ rn) utm_with_two_args n" using utm_with_two_args_halt_eq[of tp i "args" stp m rs k rn] assms apply(auto simp: bin_wc_eq tape_of_list_def tape_of_nat_def) apply(rename_tac stpa) apply(rule_tac x = stpa in exI, simp) done qed qed thus "?thesis" apply(auto simp: Hoare_halt_def utm_pre_tm_def) apply(case_tac "steps0 (Suc 0, [], ) (wcode_tm |+| utm_with_two_args) n",simp) by auto qed subsection \The correctness of utm, the halt case\ lemma utm_halt_lemma': assumes composable_tm: "composable_tm (tp, 0)" and result: "0 < rs" and args: "args \ []" and exec: "steps0 (Suc 0, Bk\(i), ) tp stp = (0, Bk\(m), Oc\(rs)@Bk\(k))" shows "\stp m n. steps0 (Suc 0, [], ) utm stp = (0, Bk\(m), Oc\(rs) @ Bk\(n))" using utm_halt_lemma_pre[of tp rs args i stp m k] assms apply(simp add: utm_pre_tm_def utm_with_two_args_def utm_def f_aprog_def f_tprog_tm_def) apply(case_tac "rec_ci rec_F", simp) done definition TSTD:: "config \ bool" where "TSTD c = (let (st, l, r) = c in st = 0 \ (\ m. l = Bk\(m)) \ (\ rs n. r = Oc\(Suc rs) @ Bk\(n)))" lemma nstd_case1: "0 < a \ NSTD (trpl_code (a, b, c))" by(simp add: NSTD.simps trpl_code.simps) lemma nonzero_bl2wc[simp]: "\m. b \ Bk\(m) \ 0 < bl2wc b" proof - have "\m. b \ Bk \ m \ bl2wc b = 0 \ False" proof(induct b) case (Cons a b) then show ?case apply(simp add: bl2wc.simps, case_tac a, simp_all add: bl2nat.simps bl2nat_double) apply(case_tac "\ m. b = Bk\(m)", erule exE) apply(metis append_Nil2 replicate_Suc_iff_anywhere) by simp qed auto thus "\m. b \ Bk\(m) \ 0 < bl2wc b" by auto qed lemma nstd_case2: "\m. b \ Bk\(m) \ NSTD (trpl_code (a, b, c))" apply(simp add: NSTD.simps trpl_code.simps) done lemma even_not_odd[elim]: "Suc (2 * x) = 2 * y \ RR" proof(induct x arbitrary: y) case (Suc x) thus ?case by(cases y;auto) qed auto declare replicate_Suc[simp del] lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\n. c = Bk\(n))" proof(induct c) case (Cons a c) then show ?case by (cases a;auto simp: bl2nat.simps bl2nat_double Cons_replicate_eq) qed (auto simp: bl2nat.simps) lemma bl2wc_exp_ex: "\Suc (bl2wc c) = 2 ^ m\ \ \ rs n. c = Oc\(rs) @ Bk\(n)" proof(induct c arbitrary: m) case (Cons a c m) { fix n have "Bk # Bk \ n = Oc \ 0 @ Bk \ Suc n" by (auto simp:replicate_Suc) hence "\rs na. Bk # Bk \ n = Oc \ rs @ Bk \ na" by blast } with Cons show ?case apply(cases a, auto) apply(case_tac m, simp_all add: bl2wc.simps, auto) apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double Cons) apply(case_tac m, simp, simp add: bin_wc_eq bl2wc.simps twice_power ) by (metis Cons.hyps Suc_pred bl2wc.simps neq0_conv power_not_zero replicate_Suc_iff_anywhere zero_neq_numeral) qed (simp add: bl2wc.simps bl2nat.simps) lemma lg_bin: assumes "\rs n. c \ Oc\(Suc rs) @ Bk\(n)" "bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0" shows "bl2wc c = 0" proof - from assms obtain rs nat n where *:"2 ^ rs - Suc 0 = nat" "c = Oc \ rs @ Bk \ n" using bl2wc_exp_ex[of c "lg (Suc (bl2wc c)) 2"] by(case_tac "(2::nat) ^ lg (Suc (bl2wc c)) 2", simp, simp, erule_tac exE, erule_tac exE, simp) have r:"bl2wc (Oc \ rs) = nat" by (metis "*"(1) bl2nat_exp_zero bl2wc.elims) hence "Suc (bl2wc c) = 2^rs" using * by(case_tac "(2::nat)^rs", auto) thus ?thesis using * assms(1) apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE) by(case_tac rs, simp, simp) qed lemma nstd_case3: "\rs n. c \ Oc\(Suc rs) @ Bk\(n) \ NSTD (trpl_code (a, b, c))" apply(simp add: NSTD.simps trpl_code.simps) apply(auto) apply(drule_tac lg_bin, simp_all) done lemma NSTD_1: "\ TSTD (a, b, c) \ rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0" using NSTD_lemma1[of "trpl_code (a, b, c)"] NSTD_lemma2[of "trpl_code (a, b, c)"] apply(simp add: TSTD_def) apply(erule_tac disjE, erule_tac nstd_case1) apply(erule_tac disjE, erule_tac nstd_case2) apply(erule_tac nstd_case3) done lemma nonstop_t_uhalt_eq: "\composable_tm (tp, 0); steps0 (Suc 0, Bk\(l), ) tp stp = (a, b, c); \ TSTD (a, b, c)\ \ rec_exec rec_nonstop [code tp, bl2wc (), stp] = Suc 0" apply(simp add: rec_nonstop_def rec_exec.simps) apply(subgoal_tac "rec_exec rec_conf [code tp, bl2wc (), stp] = trpl_code (a, b, c)", simp) apply(erule_tac NSTD_1) using rec_t_eq_steps[of tp l lm stp] apply(simp) done lemma nonstop_true: "\composable_tm (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ \ \y. rec_exec rec_nonstop ([code tp, bl2wc (), y]) = (Suc 0)" proof fix y assume a:"composable_tm0 tp" "\stp. \ TSTD (steps0 (Suc 0, Bk \ l, ) tp stp)" hence "\ TSTD (steps0 (Suc 0, Bk \ l, ) tp y)" by auto thus "rec_exec rec_nonstop [code tp, bl2wc (), y] = Suc 0" by (cases "steps0 (Suc 0, Bk\(l), ) tp y") (auto intro: nonstop_t_uhalt_eq[OF a(1)]) qed lemma cn_arity: "rec_ci (Cn n f gs) = (a, b, c) \ b = n" by(case_tac "rec_ci f", simp add: rec_ci.simps) lemma mn_arity: "rec_ci (Mn n f) = (a, b, c) \ b = n" by(case_tac "rec_ci f", simp add: rec_ci.simps) lemma f_aprog_uhalt: assumes "composable_tm (tp,0)" and unhalt: "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" and compile: "rec_ci rec_F = (F_ap, rs_pos, a_md)" shows "\\ nl. nl = [code tp, bl2wc ()] @ 0\(a_md - rs_pos ) @ suflm\ (F_ap) \" using compile proof(simp only: rec_F_def) assume h: "rec_ci (Cn (Suc (Suc 0)) rec_valu [Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]]]) = (F_ap, rs_pos, a_md)" moreover hence "rs_pos = Suc (Suc 0)" using cn_arity by simp moreover obtain ap1 ar1 ft1 where a: "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]]) = (ap1, ar1, ft1)" by(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]])", auto) moreover hence b: "ar1 = Suc (Suc 0)" using cn_arity by simp ultimately show "?thesis" proof(rule_tac i = 0 in cn_unhalt_case, auto) fix anything obtain ap2 ar2 ft2 where c: "rec_ci (Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]) = (ap2, ar2, ft2)" by(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt])", auto) moreover hence d:"ar2 = Suc (Suc 0)" using cn_arity by simp ultimately have "\\nl. nl = [code tp, bl2wc ()] @ 0 \ (ft1 - Suc (Suc 0)) @ anything\ ap1 \" using a b c d proof(rule_tac i = 0 in cn_unhalt_case, auto) fix anything obtain ap3 ar3 ft3 where e: "rec_ci rec_halt = (ap3, ar3, ft3)" by(case_tac "rec_ci rec_halt", auto) hence f: "ar3 = Suc (Suc 0)" using mn_arity by(simp add: rec_halt_def) have "\\nl. nl = [code tp, bl2wc ()] @ 0 \ (ft2 - Suc (Suc 0)) @ anything\ ap2 \" using c d e f proof(rule_tac i = 2 in cn_unhalt_case, auto simp: rec_halt_def) fix anything have "\\nl. nl = [code tp, bl2wc ()] @ 0 \ (ft3 - Suc (Suc 0)) @ anything\ ap3 \" using e f proof(rule_tac mn_unhalt_case, auto simp: rec_halt_def) fix i show "terminate rec_nonstop [code tp, bl2wc (), i]" by(rule_tac primerec_terminate, auto) next fix i show "0 < rec_exec rec_nonstop [code tp, bl2wc (), i]" using assms by(drule_tac nonstop_true, auto) qed thus "\\nl. nl = code tp # bl2wc () # 0 \ (ft3 - Suc (Suc 0)) @ anything\ ap3 \" by simp next fix apj arj ftj j anything assume "j<2" "rec_ci ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) = (apj, arj, ftj)" hence "\\nl. nl = [code tp, bl2wc ()] @ 0 \ (ftj - arj) @ anything\ apj \\nl. nl = [code tp, bl2wc ()] @ rec_exec ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) [code tp, bl2wc ()] # 0 \ (ftj - Suc arj) @ anything\" apply(rule_tac recursive_compile_correct) apply(case_tac j, auto) apply(rule_tac [!] primerec_terminate) by(auto) thus "\\nl. nl = code tp # bl2wc () # 0 \ (ftj - arj) @ anything\ apj \\nl. nl = code tp # bl2wc () # rec_exec ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) [code tp, bl2wc ()] # 0 \ (ftj - Suc arj) @ anything\" by simp next fix j assume "(j::nat) < 2" thus "terminate ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) [code tp, bl2wc ()]" by(case_tac j, auto intro!: primerec_terminate) qed thus "\\nl. nl = code tp # bl2wc () # 0 \ (ft2 - Suc (Suc 0)) @ anything\ ap2 \" by simp qed thus "\\nl. nl = code tp # bl2wc () # 0 \ (ft1 - Suc (Suc 0)) @ anything\ ap1 \" by simp qed qed lemma uabc_uhalt': "\composable_tm (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp)); rec_ci rec_F = (ap, pos, md)\ \ \\ nl. nl = [code tp, bl2wc ()]\ ap \" proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md and suflm = "[]" in f_aprog_uhalt, auto simp: abc_Hoare_unhalt_def, case_tac "abc_steps_l (0, [code tp, bl2wc ()]) ap n", simp) fix n a b assume h: "\n. abc_notfinal (abc_steps_l (0, code tp # bl2wc () # 0 \ (md - pos)) ap n) ap" "abc_steps_l (0, [code tp, bl2wc ()]) ap n = (a, b)" "composable_tm (tp, 0)" "rec_ci rec_F = (ap, pos, md)" moreover have a: "ap \ []" using h rec_ci_not_null[of "rec_F" pos md] by auto ultimately show "a < length ap" proof(erule_tac x = n in allE) assume g: "abc_notfinal (abc_steps_l (0, code tp # bl2wc () # 0 \ (md - pos)) ap n) ap" obtain ss nl where b : "abc_steps_l (0, code tp # bl2wc () # 0 \ (md - pos)) ap n = (ss, nl)" by (metis prod.exhaust) then have c: "ss < length ap" using g by simp thus "?thesis" using a b c using abc_list_crsp_steps[of "[code tp, bl2wc ()]" "md - pos" ap n ss nl] h by(simp) qed qed lemma uabc_uhalt: "\composable_tm (tp, 0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ \ \\ nl. nl = [code tp, bl2wc ()]\ f_aprog \ " proof - obtain a b c where abc:"rec_ci rec_F = (a,b,c)" by (cases "rec_ci rec_F") force assume a:"composable_tm (tp, 0)" "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" from uabc_uhalt'[OF a abc] abc_Hoare_plus_unhalt1 show "\\ nl. nl = [code tp, bl2wc ()]\ f_aprog \" by(simp add: f_aprog_def abc) qed lemma tutm_uhalt': assumes composable_tm: "composable_tm (tp,0)" and unhalt: "\ stp. (\ TSTD (steps0 (1, Bk\(l), ) tp stp))" shows "\ stp. \ is_final (steps0 (1, [Bk, Bk], <[code tp, bl2wc ()]>) utm_with_two_args stp)" unfolding utm_with_two_args_def proof(rule_tac compile_correct_unhalt, auto) show "f_tprog_tm = tm_of f_aprog" by(simp add: f_tprog_tm_def) next show "crsp (layout_of f_aprog) (0, [code tp, bl2wc ()]) (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) []" by(auto simp: crsp.simps start_of.simps) next fix stp a b show "abc_steps_l (0, [code tp, bl2wc ()]) f_aprog stp = (a, b) \ a < length f_aprog" using assms apply(drule_tac uabc_uhalt, auto simp: abc_Hoare_unhalt_def) by(erule_tac x = stp in allE, erule_tac x = stp in allE, simp) qed lemma tinres_commute: "tinres r r' \ tinres r' r" apply(auto simp: tinres_def) done lemma inres_tape: "\steps0 (st, l, r) tp stp = (a, b, c); steps0 (st, l', r') tp stp = (a', b', c'); tinres l l'; tinres r r'\ \ a = a' \ tinres b b' \ tinres c c'" proof(case_tac "steps0 (st, l', r) tp stp") fix aa ba ca assume h: "steps0 (st, l, r) tp stp = (a, b, c)" "steps0 (st, l', r') tp stp = (a', b', c')" "tinres l l'" "tinres r r'" "steps0 (st, l', r) tp stp = (aa, ba, ca)" have "tinres b ba \ c = ca \ a = aa" using h apply(rule_tac tinres_steps1, auto) done moreover have "b' = ba \ tinres c' ca \ a' = aa" using h apply(rule_tac tinres_steps2, auto intro: tinres_commute) done ultimately show "?thesis" apply(auto intro: tinres_commute) done qed lemma tape_normalize: assumes "\ stp. \ is_final(steps0 (Suc 0, [Bk,Bk], <[code tp, bl2wc ()]>) utm_with_two_args stp)" shows "\ stp. \ is_final (steps0 (Suc 0, Bk\(m), <[code tp, bl2wc ()]> @ Bk\(n)) utm_with_two_args stp)" (is "\ stp. ?P stp") proof fix stp from assms[rule_format,of stp] show "?P stp" apply(case_tac "steps0 (Suc 0, Bk\(m), <[code tp, bl2wc ()]> @ Bk\(n)) utm_with_two_args stp", simp) apply(case_tac "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc ()]>) utm_with_two_args stp", simp) apply(drule_tac inres_tape, auto) apply(auto simp: tinres_def) apply(case_tac "m > Suc (Suc 0)") apply(rule_tac x = "m - Suc (Suc 0)" in exI) apply(case_tac m, simp_all) apply(metis Suc_lessD Suc_pred replicate_Suc) apply(rule_tac x = "2 - m" in exI, simp add: replicate_add[THEN sym]) apply(simp only: numeral_2_eq_2, simp add: replicate_Suc) done qed lemma tutm_uhalt: "\composable_tm (tp,0); \ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))\ \ \ stp. \ is_final (steps0 (Suc 0, Bk\(m), <[code tp, bl2wc ()]> @ Bk\(n)) utm_with_two_args stp)" apply(rule_tac tape_normalize) apply(rule_tac tutm_uhalt'[simplified], simp_all) done lemma utm_uhalt_lemma_pre: assumes composable_tm: "composable_tm (tp, 0)" and exec: "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" and args: "args \ []" shows "\ stp. \ is_final (steps0 (Suc 0, [], ) utm_pre_tm stp)" proof - let ?P1 = "\ (l, r). l = [] \ r = " let ?Q1 = "\ (l, r). (l = [Bk] \ (\ rn. r = Oc\(Suc (code tp)) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn)))" let ?P2 = ?Q1 have "\?P1\ (wcode_tm |+| utm_with_two_args) \" proof(rule_tac Hoare_plus_unhalt) show "composable_tm (wcode_tm, 0)" by auto next show "\?P1\ wcode_tm \?Q1\" apply(rule_tac Hoare_haltI, auto) using wcode_lemma_1[of args "code tp"] args apply(auto) by (metis (mono_tags, lifting) holds_for.simps is_finalI old.prod.case) next show "\?P2\ utm_with_two_args \" proof(rule_tac Hoare_unhaltI, auto) fix n rn assume h: "is_final (steps0 (Suc 0, [Bk], Oc \ Suc (code tp) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn) utm_with_two_args n)" have "\ stp. \ is_final (steps0 (Suc 0, Bk\(Suc 0), <[code tp, bl2wc ()]> @ Bk\(rn)) utm_with_two_args stp)" using assms apply(rule_tac tutm_uhalt, simp_all) done thus "False" using h apply(erule_tac x = n in allE) apply(simp add: tape_of_list_def bin_wc_eq tape_of_nat_def) done qed qed thus "?thesis" apply(simp add: Hoare_unhalt_def utm_pre_tm_def) done qed subsection \The correctness of utm, the unhalt case.\ lemma utm_uhalt_lemma': assumes composable_tm: "composable_tm (tp, 0)" and unhalt: "\ stp. (\ TSTD (steps0 (Suc 0, Bk\(l), ) tp stp))" and args: "args \ []" shows " \ stp. \ is_final (steps0 (Suc 0, [], ) utm stp)" using utm_uhalt_lemma_pre[of tp l args] assms apply(simp add: utm_pre_tm_def utm_with_two_args_def utm_def f_aprog_def f_tprog_tm_def) apply(case_tac "rec_ci rec_F", simp) done lemma utm_halt_lemma: assumes composable_tm: "composable_tm (p, 0)" and result: "rs > 0" and args: "(args::nat list) \ []" and exec: "\(\tp. tp = (Bk\i, ))\ p \(\tp. tp = (Bk\m, Oc\rs @ Bk\k))\" shows "\(\tp. tp = ([], ))\ utm \(\tp. (\ m n. tp = (Bk\m, Oc\rs @ Bk\n)))\" proof - let ?steps0 = "steps0 (Suc 0, [], )" let ?stepsBk = "steps0 (Suc 0, Bk\i, ) p" from wcode_lemma_1[OF args,of "code p"] obtain stp ln rn where wcl1:"?steps0 wcode_tm stp = (0, [Bk], Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn)" by fast from exec Hoare_halt_def obtain n where n:"\\tp. tp = (Bk \ i, )\ p \\tp. tp = (Bk \ m, Oc \ rs @ Bk \ k)\" "is_final (?stepsBk n)" "(\tp. tp = (Bk \ m, Oc \ rs @ Bk \ k)) holds_for steps0 (Suc 0, Bk \ i, ) p n" by auto obtain a where a:"a = fst (rec_ci rec_F)" by blast have "\(\ (l, r). l = [] \ r = ) \ (wcode_tm |+| utm_with_two_args) \(\ (l, r). (\ m. l = Bk\m) \ (\ n. r = Oc\rs @ Bk\n))\" proof(rule_tac Hoare_plus_halt) show "\\(l, r). l = [] \ r = \ wcode_tm \\ (l, r). (l = [Bk] \ (\ rn. r = Oc\(Suc (code p)) @ Bk # Oc\(Suc (bl_bin ())) @ Bk\(rn)))\" using wcl1 by (auto intro!:Hoare_haltI exI[of _ stp]) next have "\ stp. (?stepsBk stp = (0, Bk\m, Oc\rs @ Bk\k))" using n by (case_tac "?stepsBk n", auto) then obtain stp where k: "steps0 (Suc 0, Bk\i, ) p stp = (0, Bk\m, Oc\rs @ Bk\k)" .. thus "\\(l, r). l = [Bk] \ (\rn. r = Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn)\ utm_with_two_args \\(l, r). (\m. l = Bk \ m) \ (\n. r = Oc \ rs @ Bk \ n)\" proof(rule_tac Hoare_haltI, auto) fix rn from utm_with_two_args_halt_eq[OF assms(1) k assms(2),of rn] assms k have "\ ma n stp. steps0 (Suc 0, [Bk], <[code p, bl2wc ()]> @ Bk \ rn) utm_with_two_args stp = (0, Bk \ ma, Oc \ rs @ Bk \ n)" by (auto simp add: bin_wc_eq) then obtain stpx m' n' where t:"steps0 (Suc 0, [Bk], <[code p, bl2wc ()]> @ Bk \ rn) utm_with_two_args stpx = (0, Bk \ m', Oc \ rs @ Bk \ n')" by auto show "\n. is_final (steps0 (Suc 0, [Bk], Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn) utm_with_two_args n) \ (\(l, r). (\m. l = Bk \ m) \ (\n. r = Oc \ rs @ Bk \ n)) holds_for steps0 (Suc 0, [Bk], Oc \ Suc (code p) @ Bk # Oc \ Suc (bl_bin ()) @ Bk \ rn) utm_with_two_args n" using t by(auto simp: bin_wc_eq tape_of_list_def tape_of_nat_def intro:exI[of _ stpx]) qed next show "composable_tm0 wcode_tm" by auto qed then obtain n where "is_final (?steps0 (wcode_tm |+| utm_with_two_args) n)" "(\(l, r). (\m. l = Bk \ m) \ (\n. r = Oc \ rs @ Bk \ n)) holds_for ?steps0 (wcode_tm |+| utm_with_two_args) n" by(auto simp add: Hoare_halt_def a) thus "?thesis" apply(case_tac "rec_ci rec_F") apply(auto simp add: utm_def Hoare_halt_def) apply(case_tac "(?steps0 (wcode_tm |+| utm_with_two_args) n)") apply(rule_tac x="n" in exI) apply(auto simp add:a utm_with_two_args_def f_aprog_def f_tprog_tm_def) done qed lemma utm_halt_lemma2: assumes composable_tm: "composable_tm (p, 0)" and args: "(args::nat list) \ []" and exec: "\(\tp. tp = ([], ))\ p \(\tp. tp = (Bk\m, <(n::nat)> @ Bk\k))\" shows "\(\tp. tp = ([], ))\ utm \(\tp. (\ m k. tp = (Bk\m, @ Bk\k)))\" using utm_halt_lemma[OF assms(1) _ assms(2), where i="0"] using assms(3) by(simp add: tape_of_nat_def) lemma utm_unhalt_lemma: assumes composable_tm: "composable_tm (p, 0)" and unhalt: "\(\tp. tp = (Bk\i, ))\ p \" and args: "args \ []" shows "\(\tp. tp = ([], ))\ utm \" proof - have "(\ TSTD (steps0 (Suc 0, Bk\(i), ) p stp))" for stp (* in unhalt, we substitute inner 'forall' n\stp *) using unhalt[unfolded Hoare_unhalt_def,rule_format,OF refl,of stp] by(cases "steps0 (Suc 0, Bk \ i, ) p stp",auto simp: Hoare_unhalt_def TSTD_def) then have "\ stp. \ is_final (steps0 (Suc 0, [], ) utm stp)" using assms by(intro utm_uhalt_lemma', auto) thus "?thesis" by(simp add: Hoare_unhalt_def) qed lemma utm_unhalt_lemma2: assumes "composable_tm (p, 0)" and "\(\tp. tp = ([], ))\ p \" and "args \ []" shows "\(\tp. tp = ([], ))\ utm \" using utm_unhalt_lemma[OF assms(1), where i="0"] using assms(2-3) by(simp add: tape_of_nat_def) end diff --git a/thys/Universal_Turing_Machine/WeakCopyTM.thy b/thys/Universal_Turing_Machine/WeakCopyTM.thy --- a/thys/Universal_Turing_Machine/WeakCopyTM.thy +++ b/thys/Universal_Turing_Machine/WeakCopyTM.thy @@ -1,1179 +1,1178 @@ (* Title: thys/WeakCopyTM.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten Cleanup of proofs Author: Franz Regensburger (FABR) 02/2022 *) subsection \Machines that duplicate a single Numeral\ subsubsection \A Turing machine that duplicates its input if the input is a single numeral\ text\The Machine WeakCopyTM does not check the number of its arguments on the initial tape. If it is provided a single numeral it does a perfect job. However, if it gets no or more than one argument, it does not complain but produces some result.\ theory WeakCopyTM imports Turing_HaltingConditions begin declare adjust.simps[simp del] (* ---------- tm_copy_begin_orig ---------- *) (* let tm_copy_begin_orig = from_raw [ (WB,0),(R,2), (R,3),(R,2), (WO,3),(L,4), (L,4),(L,0) ] *) definition tm_copy_begin_orig :: "instr list" where "tm_copy_begin_orig \ [(WB,0),(R,2), (R,3),(R,2), (WO,3),(L,4), (L,4),(L,0)]" fun inv_begin0 :: "nat \ tape \ bool" and inv_begin1 :: "nat \ tape \ bool" and inv_begin2 :: "nat \ tape \ bool" and inv_begin3 :: "nat \ tape \ bool" and inv_begin4 :: "nat \ tape \ bool" where "inv_begin0 n (l, r) = ((n > 1 \ (l, r) = (Oc \ (n - 2), [Oc, Oc, Bk, Oc])) \ (n = 1 \ (l, r) = ([], [Bk, Oc, Bk, Oc])))" | "inv_begin1 n (l, r) = ((l, r) = ([], Oc \ n))" | "inv_begin2 n (l, r) = (\ i j. i > 0 \ i + j = n \ (l, r) = (Oc \ i, Oc \ j))" | "inv_begin3 n (l, r) = (n > 0 \ (l, tl r) = (Bk # Oc \ n, []))" | "inv_begin4 n (l, r) = (n > 0 \ (l, r) = (Oc \ n, [Bk, Oc]) \ (l, r) = (Oc \ (n - 1), [Oc, Bk, Oc]))" fun inv_begin :: "nat \ config \ bool" where "inv_begin n (s, tap) = (if s = 0 then inv_begin0 n tap else if s = 1 then inv_begin1 n tap else if s = 2 then inv_begin2 n tap else if s = 3 then inv_begin3 n tap else if s = 4 then inv_begin4 n tap else False)" lemma inv_begin_step_E: "\0 < i; 0 < j\ \ \ia>0. ia + j - Suc 0 = i + j \ Oc # Oc \ i = Oc \ ia" by (rule_tac x = "Suc i" in exI, simp) lemma inv_begin_step: assumes "inv_begin n cf" and "n > 0" shows "inv_begin n (step0 cf tm_copy_begin_orig)" using assms unfolding tm_copy_begin_orig_def apply(cases cf) apply(auto simp: numeral_eqs_upto_12 split: if_splits elim:inv_begin_step_E) apply(cases "hd (snd (snd cf))";cases "(snd (snd cf))",auto) done lemma inv_begin_steps: assumes "inv_begin n cf" and "n > 0" shows "inv_begin n (steps0 cf tm_copy_begin_orig stp)" apply(induct stp) apply(simp add: assms) apply(auto simp del: steps.simps) apply(rule_tac inv_begin_step) apply(simp_all add: assms) done lemma begin_partial_correctness: assumes "is_final (steps0 (1, [], Oc \ n) tm_copy_begin_orig stp)" shows "0 < n \ \inv_begin1 n\ tm_copy_begin_orig \inv_begin0 n\" proof(rule_tac Hoare_haltI) fix l r assume h: "0 < n" "inv_begin1 n (l, r)" have "inv_begin n (steps0 (1, [], Oc \ n) tm_copy_begin_orig stp)" using h by (rule_tac inv_begin_steps) (simp_all) then show "\stp. is_final (steps0 (1, l, r) tm_copy_begin_orig stp) \ inv_begin0 n holds_for steps (1, l, r) (tm_copy_begin_orig, 0) stp" using h assms apply(rule_tac x = stp in exI) apply(case_tac "(steps0 (1, [], Oc \ n) tm_copy_begin_orig stp)", simp) done qed fun measure_begin_state :: "config \ nat" where "measure_begin_state (s, l, r) = (if s = 0 then 0 else 5 - s)" fun measure_begin_step :: "config \ nat" where "measure_begin_step (s, l, r) = (if s = 2 then length r else if s = 3 then (if r = [] \ r = [Bk] then 1 else 0) else if s = 4 then length l else 0)" definition "measure_begin = measures [measure_begin_state, measure_begin_step]" lemma wf_measure_begin: shows "wf measure_begin" unfolding measure_begin_def by auto lemma measure_begin_induct [case_names Step]: "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_begin\ \ \n. P (f n)" using wf_measure_begin by (metis wf_iff_no_infinite_down_chain) lemma begin_halts: assumes h: "x > 0" shows "\ stp. is_final (steps0 (1, [], Oc \ x) tm_copy_begin_orig stp)" proof (induct rule: measure_begin_induct) case (Step n) have "\ is_final (steps0 (1, [], Oc \ x) tm_copy_begin_orig n)" by fact moreover have "inv_begin x (steps0 (1, [], Oc \ x) tm_copy_begin_orig n)" by (rule_tac inv_begin_steps) (simp_all add: h) moreover obtain s l r where eq: "(steps0 (1, [], Oc \ x) tm_copy_begin_orig n) = (s, l, r)" by (metis measure_begin_state.cases) ultimately have "(step0 (s, l, r) tm_copy_begin_orig, s, l, r) \ measure_begin" apply(auto simp: measure_begin_def tm_copy_begin_orig_def numeral_eqs_upto_12 split: if_splits) apply(subgoal_tac "r = [Oc]") apply(auto) by (metis cell.exhaust list.exhaust list.sel(3)) then show "(steps0 (1, [], Oc \ x) tm_copy_begin_orig (Suc n), steps0 (1, [], Oc \ x) tm_copy_begin_orig n) \ measure_begin" using eq by (simp only: step_red) qed lemma begin_correct: shows "0 < n \ \inv_begin1 n\ tm_copy_begin_orig \inv_begin0 n\" using begin_partial_correctness begin_halts by blast lemma begin_correct2: assumes "0 < (n::nat)" shows "\\tap. tap = ([]::cell list, Oc \ n)\ tm_copy_begin_orig \\tap. (n > 1 \ tap = (Oc \ (n - 2), [Oc, Oc, Bk, Oc])) \ (n = 1 \ tap = ([]::cell list, [Bk, Oc, Bk, Oc])) \" proof - from assms have "\inv_begin1 n\ tm_copy_begin_orig \inv_begin0 n\" using begin_partial_correctness begin_halts by blast with assms have "\\tap. tap = ([]::cell list, Oc \ n)\ tm_copy_begin_orig \inv_begin0 n\" using Hoare_haltE Hoare_haltI inv_begin1.simps by presburger with assms show ?thesis - by (smt Hoare_haltI Hoare_halt_def Pair_inject + by (smt (verit) Hoare_haltI Hoare_halt_def Pair_inject holds_for.elims(2) holds_for.simps inv_begin0.simps is_final.elims(2)) qed (* ---------- tm_copy_loop_orig ---------- *) (* Delete some theorems from the simpset *) declare seq_tm.simps [simp del] declare shift.simps[simp del] declare composable_tm.simps[simp del] declare step.simps[simp del] declare steps.simps[simp del] (* let tm_copy_loop_orig = from_raw [ (R,0),(R,2), (R,3),(WB,2), (R,3),(R,4), (WO,5),(R,4), (L,6),(L,5), (L,6),(L,1) ] *) definition tm_copy_loop_orig :: "instr list" where "tm_copy_loop_orig \ [ (R, 0),(R, 2), (R, 3),(WB, 2), (R, 3),(R, 4), (WO, 5),(R, 4), (L, 6),(L, 5), (L, 6),(L, 1) ]" fun inv_loop1_loop :: "nat \ tape \ bool" and inv_loop1_exit :: "nat \ tape \ bool" and inv_loop5_loop :: "nat \ tape \ bool" and inv_loop5_exit :: "nat \ tape \ bool" and inv_loop6_loop :: "nat \ tape \ bool" and inv_loop6_exit :: "nat \ tape \ bool" where "inv_loop1_loop n (l, r) = (\ i j. i + j + 1 = n \ (l, r) = (Oc\i, Oc#Oc#Bk\j @ Oc\j) \ j > 0)" | "inv_loop1_exit n (l, r) = (0 < n \ (l, r) = ([], Bk#Oc#Bk\n @ Oc\n))" | "inv_loop5_loop x (l, r) = (\ i j k t. i + j = Suc x \ i > 0 \ j > 0 \ k + t = j \ t > 0 \ (l, r) = (Oc\k@Bk\j@Oc\i, Oc\t))" | "inv_loop5_exit x (l, r) = (\ i j. i + j = Suc x \ i > 0 \ j > 0 \ (l, r) = (Bk\(j - 1)@Oc\i, Bk # Oc\j))" | "inv_loop6_loop x (l, r) = (\ i j k t. i + j = Suc x \ i > 0 \ k + t + 1 = j \ (l, r) = (Bk\k @ Oc\i, Bk\(Suc t) @ Oc\j))" | "inv_loop6_exit x (l, r) = (\ i j. i + j = x \ j > 0 \ (l, r) = (Oc\i, Oc#Bk\j @ Oc\j))" fun inv_loop0 :: "nat \ tape \ bool" and inv_loop1 :: "nat \ tape \ bool" and inv_loop2 :: "nat \ tape \ bool" and inv_loop3 :: "nat \ tape \ bool" and inv_loop4 :: "nat \ tape \ bool" and inv_loop5 :: "nat \ tape \ bool" and inv_loop6 :: "nat \ tape \ bool" where "inv_loop0 n (l, r) = (0 < n \ (l, r) = ([Bk], Oc # Bk\n @ Oc\n))" | "inv_loop1 n (l, r) = (inv_loop1_loop n (l, r) \ inv_loop1_exit n (l, r))" | "inv_loop2 n (l, r) = (\ i j any. i + j = n \ n > 0 \ i > 0 \ j > 0 \ (l, r) = (Oc\i, any#Bk\j@Oc\j))" | "inv_loop3 n (l, r) = (\ i j k t. i + j = n \ i > 0 \ j > 0 \ k + t = Suc j \ (l, r) = (Bk\k@Oc\i, Bk\t@Oc\j))" | "inv_loop4 n (l, r) = (\ i j k t. i + j = n \ i > 0 \ j > 0 \ k + t = j \ (l, r) = (Oc\k @ Bk\(Suc j)@Oc\i, Oc\t))" | "inv_loop5 n (l, r) = (inv_loop5_loop n (l, r) \ inv_loop5_exit n (l, r))" | "inv_loop6 n (l, r) = (inv_loop6_loop n (l, r) \ inv_loop6_exit n (l, r))" fun inv_loop :: "nat \ config \ bool" where "inv_loop x (s, l, r) = (if s = 0 then inv_loop0 x (l, r) else if s = 1 then inv_loop1 x (l, r) else if s = 2 then inv_loop2 x (l, r) else if s = 3 then inv_loop3 x (l, r) else if s = 4 then inv_loop4 x (l, r) else if s = 5 then inv_loop5 x (l, r) else if s = 6 then inv_loop6 x (l, r) else False)" declare inv_loop.simps[simp del] inv_loop1.simps[simp del] inv_loop2.simps[simp del] inv_loop3.simps[simp del] inv_loop4.simps[simp del] inv_loop5.simps[simp del] inv_loop6.simps[simp del] lemma inv_loop3_Bk_empty_via_2[elim]: "\0 < x; inv_loop2 x (b, [])\ \ inv_loop3 x (Bk # b, [])" by (auto simp: inv_loop2.simps inv_loop3.simps) lemma inv_loop3_Bk_empty[elim]: "\0 < x; inv_loop3 x (b, [])\ \ inv_loop3 x (Bk # b, [])" by (auto simp: inv_loop3.simps) lemma inv_loop5_Oc_empty_via_4[elim]: "\0 < x; inv_loop4 x (b, [])\ \ inv_loop5 x (b, [Oc])" by(auto simp: inv_loop4.simps inv_loop5.simps;force) lemma inv_loop1_Bk[elim]: "\0 < x; inv_loop1 x (b, Bk # list)\ \ list = Oc # Bk \ x @ Oc \ x" by (auto simp: inv_loop1.simps) lemma inv_loop3_Bk_via_2[elim]: "\0 < x; inv_loop2 x (b, Bk # list)\ \ inv_loop3 x (Bk # b, list)" by(auto simp: inv_loop2.simps inv_loop3.simps;force) lemma inv_loop3_Bk_move[elim]: "\0 < x; inv_loop3 x (b, Bk # list)\ \ inv_loop3 x (Bk # b, list)" apply(auto simp: inv_loop3.simps) apply (rename_tac i j k t) apply(rule_tac [!] x = i in exI, rule_tac [!] x = j in exI, simp_all) apply(case_tac [!] t, auto) done lemma inv_loop5_Oc_via_4_Bk[elim]: "\0 < x; inv_loop4 x (b, Bk # list)\ \ inv_loop5 x (b, Oc # list)" by (auto simp: inv_loop4.simps inv_loop5.simps) lemma inv_loop6_Bk_via_5[elim]: "\0 < x; inv_loop5 x ([], Bk # list)\ \ inv_loop6 x ([], Bk # Bk # list)" by (auto simp: inv_loop6.simps inv_loop5.simps) lemma inv_loop5_loop_no_Bk[simp]: "inv_loop5_loop x (b, Bk # list) = False" by (auto simp: inv_loop5.simps) lemma inv_loop6_exit_no_Bk[simp]: "inv_loop6_exit x (b, Bk # list) = False" by (auto simp: inv_loop6.simps) declare inv_loop5_loop.simps[simp del] inv_loop5_exit.simps[simp del] inv_loop6_loop.simps[simp del] inv_loop6_exit.simps[simp del] lemma inv_loop6_loopBk_via_5[elim]:"\0 < x; inv_loop5_exit x (b, Bk # list); b \ []; hd b = Bk\ \ inv_loop6_loop x (tl b, Bk # Bk # list)" apply(simp only: inv_loop5_exit.simps inv_loop6_loop.simps ) apply(erule_tac exE)+ apply(rename_tac i j) apply(rule_tac x = i in exI, rule_tac x = j in exI, rule_tac x = "j - Suc (Suc 0)" in exI, rule_tac x = "Suc 0" in exI, auto) apply(case_tac [!] j, simp_all) apply(case_tac [!] "j-1", simp_all) done lemma inv_loop6_loop_no_Oc_Bk[simp]: "inv_loop6_loop x (b, Oc # Bk # list) = False" by (auto simp: inv_loop6_loop.simps) lemma inv_loop6_exit_Oc_Bk_via_5[elim]: "\x > 0; inv_loop5_exit x (b, Bk # list); b \ []; hd b = Oc\ \ inv_loop6_exit x (tl b, Oc # Bk # list)" apply(simp only: inv_loop5_exit.simps inv_loop6_exit.simps) apply(erule_tac exE)+ apply(rule_tac x = "x - 1" in exI, rule_tac x = 1 in exI, simp) apply(rename_tac i j) apply(case_tac j;case_tac "j-1", auto) done lemma inv_loop6_Bk_tail_via_5[elim]: "\0 < x; inv_loop5 x (b, Bk # list); b \ []\ \ inv_loop6 x (tl b, hd b # Bk # list)" apply(simp add: inv_loop5.simps inv_loop6.simps) apply(cases "hd b", simp_all, auto) done lemma inv_loop6_loop_Bk_Bk_drop[elim]: "\0 < x; inv_loop6_loop x (b, Bk # list); b \ []; hd b = Bk\ \ inv_loop6_loop x (tl b, Bk # Bk # list)" apply(simp only: inv_loop6_loop.simps) apply(erule_tac exE)+ apply(rename_tac i j k t) apply(rule_tac x = i in exI, rule_tac x = j in exI, rule_tac x = "k - 1" in exI, rule_tac x = "Suc t" in exI, auto) apply(case_tac [!] k, auto) done lemma inv_loop6_exit_Oc_Bk_via_loop6[elim]: "\0 < x; inv_loop6_loop x (b, Bk # list); b \ []; hd b = Oc\ \ inv_loop6_exit x (tl b, Oc # Bk # list)" apply(simp only: inv_loop6_loop.simps inv_loop6_exit.simps) apply(erule_tac exE)+ apply(rename_tac i j k t) apply(rule_tac x = "i - 1" in exI, rule_tac x = j in exI, auto) apply(case_tac [!] k, auto) done lemma inv_loop6_Bk_tail[elim]: "\0 < x; inv_loop6 x (b, Bk # list); b \ []\ \ inv_loop6 x (tl b, hd b # Bk # list)" apply(simp add: inv_loop6.simps) apply(case_tac "hd b", simp_all, auto) done lemma inv_loop2_Oc_via_1[elim]: "\0 < x; inv_loop1 x (b, Oc # list)\ \ inv_loop2 x (Oc # b, list)" apply(auto simp: inv_loop1.simps inv_loop2.simps,force) done lemma inv_loop2_Bk_via_Oc[elim]: "\0 < x; inv_loop2 x (b, Oc # list)\ \ inv_loop2 x (b, Bk # list)" by (auto simp: inv_loop2.simps) lemma inv_loop4_Oc_via_3[elim]: "\0 < x; inv_loop3 x (b, Oc # list)\ \ inv_loop4 x (Oc # b, list)" apply(auto simp: inv_loop3.simps inv_loop4.simps) apply(rename_tac i j) apply(rule_tac [!] x = i in exI, auto) apply(rule_tac [!] x = "Suc 0" in exI, rule_tac [!] x = "j - 1" in exI) apply(case_tac [!] j, auto) done lemma inv_loop4_Oc_move[elim]: assumes "0 < x" "inv_loop4 x (b, Oc # list)" shows "inv_loop4 x (Oc # b, list)" proof - from assms[unfolded inv_loop4.simps] obtain i j k t where "i + j = x" "0 < i" "0 < j" "k + t = j" "(b, Oc # list) = (Oc \ k @ Bk \ Suc j @ Oc \ i, Oc \ t)" by auto thus ?thesis unfolding inv_loop4.simps apply(rule_tac [!] x = "i" in exI,rule_tac [!] x = "j" in exI) apply(rule_tac [!] x = "Suc k" in exI,rule_tac [!] x = "t - 1" in exI) by(cases t,auto) qed lemma inv_loop5_exit_no_Oc[simp]: "inv_loop5_exit x (b, Oc # list) = False" by (auto simp: inv_loop5_exit.simps) lemma inv_loop5_exit_Bk_Oc_via_loop[elim]: " \inv_loop5_loop x (b, Oc # list); b \ []; hd b = Bk\ \ inv_loop5_exit x (tl b, Bk # Oc # list)" apply(simp only: inv_loop5_loop.simps inv_loop5_exit.simps) apply(erule_tac exE)+ apply(rename_tac i j k t) apply(rule_tac x = i in exI) apply(case_tac k, auto) done lemma inv_loop5_loop_Oc_Oc_drop[elim]: "\inv_loop5_loop x (b, Oc # list); b \ []; hd b = Oc\ \ inv_loop5_loop x (tl b, Oc # Oc # list)" apply(simp only: inv_loop5_loop.simps) apply(erule_tac exE)+ apply(rename_tac i j k t) apply(rule_tac x = i in exI, rule_tac x = j in exI) apply(rule_tac x = "k - 1" in exI, rule_tac x = "Suc t" in exI) apply(case_tac k, auto) done lemma inv_loop5_Oc_tl[elim]: "\inv_loop5 x (b, Oc # list); b \ []\ \ inv_loop5 x (tl b, hd b # Oc # list)" apply(simp add: inv_loop5.simps) apply(cases "hd b", simp_all, auto) done lemma inv_loop1_Bk_Oc_via_6[elim]: "\0 < x; inv_loop6 x ([], Oc # list)\ \ inv_loop1 x ([], Bk # Oc # list)" by(auto simp: inv_loop6.simps inv_loop1.simps inv_loop6_loop.simps inv_loop6_exit.simps) lemma inv_loop1_Oc_via_6[elim]: "\0 < x; inv_loop6 x (b, Oc # list); b \ []\ \ inv_loop1 x (tl b, hd b # Oc # list)" by(auto simp: inv_loop6.simps inv_loop1.simps inv_loop6_loop.simps inv_loop6_exit.simps) lemma inv_loop_nonempty[simp]: "inv_loop1 x (b, []) = False" "inv_loop2 x ([], b) = False" "inv_loop2 x (l', []) = False" "inv_loop3 x (b, []) = False" "inv_loop4 x ([], b) = False" "inv_loop5 x ([], list) = False" "inv_loop6 x ([], Bk # xs) = False" by (auto simp: inv_loop1.simps inv_loop2.simps inv_loop3.simps inv_loop4.simps inv_loop5.simps inv_loop6.simps inv_loop5_exit.simps inv_loop5_loop.simps inv_loop6_loop.simps) lemma inv_loop_nonemptyE[elim]: "\inv_loop5 x (b, [])\ \ RR" "inv_loop6 x (b, []) \ RR" "\inv_loop1 x (b, Bk # list)\ \ b = []" by (auto simp: inv_loop4.simps inv_loop5.simps inv_loop5_exit.simps inv_loop5_loop.simps inv_loop6.simps inv_loop6_exit.simps inv_loop6_loop.simps inv_loop1.simps) lemma inv_loop6_Bk_Bk_drop[elim]: "\inv_loop6 x ([], Bk # list)\ \ inv_loop6 x ([], Bk # Bk # list)" by (simp) lemma inv_loop_step: "\inv_loop x cf; x > 0\ \ inv_loop x (step cf (tm_copy_loop_orig, 0))" apply(cases cf, cases "snd (snd cf)"; cases "hd (snd (snd cf))") apply(auto simp: inv_loop.simps step.simps tm_copy_loop_orig_def numeral_eqs_upto_12 split: if_splits) done lemma inv_loop_steps: "\inv_loop x cf; x > 0\ \ inv_loop x (steps cf (tm_copy_loop_orig, 0) stp)" apply(induct stp, simp add: steps.simps, simp) apply(erule_tac inv_loop_step, simp) done fun loop_stage :: "config \ nat" where "loop_stage (s, l, r) = (if s = 0 then 0 else (Suc (length (takeWhile (\a. a = Oc) (rev l @ r)))))" fun loop_state :: "config \ nat" where "loop_state (s, l, r) = (if s = 2 \ hd r = Oc then 0 else if s = 1 then 1 else 10 - s)" fun loop_step :: "config \ nat" where "loop_step (s, l, r) = (if s = 3 then length r else if s = 4 then length r else if s = 5 then length l else if s = 6 then length l else 0)" definition measure_loop :: "(config \ config) set" where "measure_loop = measures [loop_stage, loop_state, loop_step]" lemma wf_measure_loop: "wf measure_loop" unfolding measure_loop_def by (auto) lemma measure_loop_induct [case_names Step]: "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_loop\ \ \n. P (f n)" using wf_measure_loop by (metis wf_iff_no_infinite_down_chain) lemma inv_loop4_not_just_Oc[elim]: "\inv_loop4 x (l', []); length (takeWhile (\a. a = Oc) (rev l' @ [Oc])) \ length (takeWhile (\a. a = Oc) (rev l'))\ \ RR" "\inv_loop4 x (l', Bk # list); length (takeWhile (\a. a = Oc) (rev l' @ Oc # list)) \ length (takeWhile (\a. a = Oc) (rev l' @ Bk # list))\ \ RR" apply(auto simp: inv_loop4.simps) apply(rename_tac i j) apply(case_tac [!] j, simp_all add: List.takeWhile_tail) done lemma takeWhile_replicate_append: "P a \ takeWhile P (a\x @ ys) = a\x @ takeWhile P ys" by (induct x, auto) lemma takeWhile_replicate: "P a \ takeWhile P (a\x) = a\x" by (induct x, auto) lemma inv_loop5_Bk_E[elim]: "\inv_loop5 x (l', Bk # list); l' \ []; length (takeWhile (\a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) \ length (takeWhile (\a. a = Oc) (rev l' @ Bk # list))\ \ RR" apply(cases "length list";cases "length list - 1" ,auto simp: inv_loop5.simps inv_loop5_exit.simps takeWhile_replicate_append takeWhile_replicate) apply(cases "length list - 2"; force simp add: List.takeWhile_tail)+ done lemma inv_loop1_hd_Oc[elim]: "\inv_loop1 x (l', Oc # list)\ \ hd list = Oc" by (auto simp: inv_loop1.simps) lemma inv_loop6_not_just_Bk[dest!]: "\length (takeWhile P (rev (tl l') @ hd l' # list)) \ length (takeWhile P (rev l' @ list))\ \ l' = []" apply(cases l', simp_all) done lemma inv_loop2_OcE[elim]: "\inv_loop2 x (l', Oc # list); l' \ []\ \ length (takeWhile (\a. a = Oc) (rev l' @ Bk # list)) < length (takeWhile (\a. a = Oc) (rev l' @ Oc # list))" apply(auto simp: inv_loop2.simps takeWhile_tail takeWhile_replicate_append takeWhile_replicate) done lemma loop_halts: assumes h: "n > 0" "inv_loop n (1, l, r)" shows "\ stp. is_final (steps0 (1, l, r) tm_copy_loop_orig stp)" proof (induct rule: measure_loop_induct) case (Step stp) have "\ is_final (steps0 (1, l, r) tm_copy_loop_orig stp)" by fact moreover have "inv_loop n (steps0 (1, l, r) tm_copy_loop_orig stp)" by (rule_tac inv_loop_steps) (simp_all only: h) moreover obtain s l' r' where eq: "(steps0 (1, l, r) tm_copy_loop_orig stp) = (s, l', r')" by (metis measure_begin_state.cases) ultimately have "(step0 (s, l', r') tm_copy_loop_orig, s, l', r') \ measure_loop" using h(1) apply(cases r';cases "hd r'") apply(auto simp: inv_loop.simps step.simps tm_copy_loop_orig_def numeral_eqs_upto_12 measure_loop_def split: if_splits) done then show "(steps0 (1, l, r) tm_copy_loop_orig (Suc stp), steps0 (1, l, r) tm_copy_loop_orig stp) \ measure_loop" using eq by (simp only: step_red) qed lemma loop_correct: assumes "0 < n" shows "\inv_loop1 n\ tm_copy_loop_orig \inv_loop0 n\" using assms proof(rule_tac Hoare_haltI) fix l r assume h: "0 < n" "inv_loop1 n (l, r)" then obtain stp where k: "is_final (steps0 (1, l, r) tm_copy_loop_orig stp)" using loop_halts apply(simp add: inv_loop.simps) apply(blast) done moreover have "inv_loop n (steps0 (1, l, r) tm_copy_loop_orig stp)" using h by (rule_tac inv_loop_steps) (simp_all add: inv_loop.simps) ultimately show "\stp. is_final (steps0 (1, l, r) tm_copy_loop_orig stp) \ inv_loop0 n holds_for steps0 (1, l, r) tm_copy_loop_orig stp" using h(1) apply(rule_tac x = stp in exI) apply(case_tac "(steps0 (1, l, r) tm_copy_loop_orig stp)") apply(simp add: inv_loop.simps) done qed (* ---------- tm_copy_end ---------- *) (* let tm_copy_end_orig = from_raw [ (L,0),(R,2), (WO,3),(L,4), (R,2),(R,2), (L,5),(WB,4), (R,0),(L,5) ] let tm_copy_end_new = from_raw [ (R,0),(R,2), -- changed action for Bk from (L,0) to (R,0) (WO,3),(L,4), (R,2),(R,2), (L,5),(WB,4), (R,0),(L,5) ] *) definition tm_copy_end_orig :: "instr list" where "tm_copy_end_orig \ [ (L, 0),(R, 2), (WO, 3),(L, 4), (R, 2),(R, 2), (L, 5),(WB, 4), (R, 0),(L, 5) ]" definition tm_copy_end_new :: "instr list" where "tm_copy_end_new \ [ (R, 0),(R, 2), (WO, 3),(L, 4), (R, 2),(R, 2), (L, 5),(WB, 4), (R, 0),(L, 5) ]" fun inv_end5_loop :: "nat \ tape \ bool" and inv_end5_exit :: "nat \ tape \ bool" where "inv_end5_loop x (l, r) = (\ i j. i + j = x \ x > 0 \ j > 0 \ l = Oc\i @ [Bk] \ r = Oc\j @ Bk # Oc\x)" | "inv_end5_exit x (l, r) = (x > 0 \ l = [] \ r = Bk # Oc\x @ Bk # Oc\x)" fun inv_end0 :: "nat \ tape \ bool" and inv_end1 :: "nat \ tape \ bool" and inv_end2 :: "nat \ tape \ bool" and inv_end3 :: "nat \ tape \ bool" and inv_end4 :: "nat \ tape \ bool" and inv_end5 :: "nat \ tape \ bool" where "inv_end0 n (l, r) = (n > 0 \ (l, r) = ([Bk], Oc\n @ Bk # Oc\n))" | "inv_end1 n (l, r) = (n > 0 \ (l, r) = ([Bk], Oc # Bk\n @ Oc\n))" | "inv_end2 n (l, r) = (\ i j. i + j = Suc n \ n > 0 \ l = Oc\i @ [Bk] \ r = Bk\j @ Oc\n)" | "inv_end3 n (l, r) = (\ i j. n > 0 \ i + j = n \ l = Oc\i @ [Bk] \ r = Oc # Bk\j@ Oc\n)" | "inv_end4 n (l, r) = (\ any. n > 0 \ l = Oc\n @ [Bk] \ r = any#Oc\n)" | "inv_end5 n (l, r) = (inv_end5_loop n (l, r) \ inv_end5_exit n (l, r))" fun inv_end :: "nat \ config \ bool" where "inv_end n (s, l, r) = (if s = 0 then inv_end0 n (l, r) else if s = 1 then inv_end1 n (l, r) else if s = 2 then inv_end2 n (l, r) else if s = 3 then inv_end3 n (l, r) else if s = 4 then inv_end4 n (l, r) else if s = 5 then inv_end5 n (l, r) else False)" declare inv_end.simps[simp del] inv_end1.simps[simp del] inv_end0.simps[simp del] inv_end2.simps[simp del] inv_end3.simps[simp del] inv_end4.simps[simp del] inv_end5.simps[simp del] lemma inv_end_nonempty[simp]: "inv_end1 x (b, []) = False" "inv_end1 x ([], list) = False" "inv_end2 x (b, []) = False" "inv_end3 x (b, []) = False" "inv_end4 x (b, []) = False" "inv_end5 x (b, []) = False" "inv_end5 x ([], Oc # list) = False" by (auto simp: inv_end1.simps inv_end2.simps inv_end3.simps inv_end4.simps inv_end5.simps) lemma inv_end0_Bk_via_1[elim]: "\0 < x; inv_end1 x (b, Bk # list); b \ []\ \ inv_end0 x (tl b, hd b # Bk # list)" by (auto simp: inv_end1.simps inv_end0.simps) lemma inv_end3_Oc_via_2[elim]: "\0 < x; inv_end2 x (b, Bk # list)\ \ inv_end3 x (b, Oc # list)" apply(auto simp: inv_end2.simps inv_end3.simps) by (metis Cons_replicate_eq One_nat_def Suc_inject Suc_pred add_Suc_right cell.distinct(1) empty_replicate list.sel(3) neq0_conv self_append_conv2 tl_append2 tl_replicate) lemma inv_end2_Bk_via_3[elim]: "\0 < x; inv_end3 x (b, Bk # list)\ \ inv_end2 x (Bk # b, list)" by (auto simp: inv_end2.simps inv_end3.simps) lemma inv_end5_Bk_via_4[elim]: "\0 < x; inv_end4 x ([], Bk # list)\ \ inv_end5 x ([], Bk # Bk # list)" by (auto simp: inv_end4.simps inv_end5.simps) lemma inv_end5_Bk_tail_via_4[elim]: "\0 < x; inv_end4 x (b, Bk # list); b \ []\ \ inv_end5 x (tl b, hd b # Bk # list)" apply(auto simp: inv_end4.simps inv_end5.simps) apply(rule_tac x = 1 in exI, simp) done lemma inv_end0_Bk_via_5[elim]: "\0 < x; inv_end5 x (b, Bk # list)\ \ inv_end0 x (Bk # b, list)" by(auto simp: inv_end5.simps inv_end0.simps gr0_conv_Suc) lemma inv_end2_Oc_via_1[elim]: "\0 < x; inv_end1 x (b, Oc # list)\ \ inv_end2 x (Oc # b, list)" by (auto simp: inv_end1.simps inv_end2.simps) lemma inv_end4_Bk_Oc_via_2[elim]: "\0 < x; inv_end2 x ([], Oc # list)\ \ inv_end4 x ([], Bk # Oc # list)" by (auto simp: inv_end2.simps inv_end4.simps) lemma inv_end4_Oc_via_2[elim]: "\0 < x; inv_end2 x (b, Oc # list); b \ []\ \ inv_end4 x (tl b, hd b # Oc # list)" by(auto simp: inv_end2.simps inv_end4.simps gr0_conv_Suc) lemma inv_end2_Oc_via_3[elim]: "\0 < x; inv_end3 x (b, Oc # list)\ \ inv_end2 x (Oc # b, list)" by (auto simp: inv_end2.simps inv_end3.simps) lemma inv_end4_Bk_via_Oc[elim]: "\0 < x; inv_end4 x (b, Oc # list)\ \ inv_end4 x (b, Bk # list)" by (auto simp: inv_end2.simps inv_end4.simps) lemma inv_end5_Bk_drop_Oc[elim]: "\0 < x; inv_end5 x ([], Oc # list)\ \ inv_end5 x ([], Bk # Oc # list)" by (auto simp: inv_end2.simps inv_end5.simps) declare inv_end5_loop.simps[simp del] inv_end5_exit.simps[simp del] lemma inv_end5_exit_no_Oc[simp]: "inv_end5_exit x (b, Oc # list) = False" by (auto simp: inv_end5_exit.simps) lemma inv_end5_loop_no_Bk_Oc[simp]: "inv_end5_loop x (tl b, Bk # Oc # list) = False" by (auto simp: inv_end5_loop.simps) lemma inv_end5_exit_Bk_Oc_via_loop[elim]: "\0 < x; inv_end5_loop x (b, Oc # list); b \ []; hd b = Bk\ \ inv_end5_exit x (tl b, Bk # Oc # list)" apply(auto simp: inv_end5_loop.simps inv_end5_exit.simps) using hd_replicate apply fastforce by (metis cell.distinct(1) hd_append2 hd_replicate list.sel(3) self_append_conv2 split_head_repeat(2)) lemma inv_end5_loop_Oc_Oc_drop[elim]: "\0 < x; inv_end5_loop x (b, Oc # list); b \ []; hd b = Oc\ \ inv_end5_loop x (tl b, Oc # Oc # list)" apply(simp only: inv_end5_loop.simps inv_end5_exit.simps) apply(erule_tac exE)+ apply(rename_tac i j) apply(rule_tac x = "i - 1" in exI, rule_tac x = "Suc j" in exI, auto) apply(case_tac [!] i, simp_all) done lemma inv_end5_Oc_tail[elim]: "\0 < x; inv_end5 x (b, Oc # list); b \ []\ \ inv_end5 x (tl b, hd b # Oc # list)" apply(simp add: inv_end2.simps inv_end5.simps) apply(case_tac "hd b", simp_all, auto) done lemma inv_end_step: "\x > 0; inv_end x cf\ \ inv_end x (step cf (tm_copy_end_new, 0))" apply(cases cf, cases "snd (snd cf)"; cases "hd (snd (snd cf))") apply(auto simp: inv_end.simps step.simps tm_copy_end_new_def numeral_eqs_upto_12 split: if_splits) apply (simp add: inv_end1.simps) done lemma inv_end_steps: "\x > 0; inv_end x cf\ \ inv_end x (steps cf (tm_copy_end_new, 0) stp)" apply(induct stp, simp add:steps.simps, simp) apply(erule_tac inv_end_step, simp) done fun end_state :: "config \ nat" where "end_state (s, l, r) = (if s = 0 then 0 else if s = 1 then 5 else if s = 2 \ s = 3 then 4 else if s = 4 then 3 else if s = 5 then 2 else 0)" fun end_stage :: "config \ nat" where "end_stage (s, l, r) = (if s = 2 \ s = 3 then (length r) else 0)" fun end_step :: "config \ nat" where "end_step (s, l, r) = (if s = 4 then (if hd r = Oc then 1 else 0) else if s = 5 then length l else if s = 2 then 1 else if s = 3 then 0 else 0)" definition end_LE :: "(config \ config) set" where "end_LE = measures [end_state, end_stage, end_step]" lemma wf_end_le: "wf end_LE" unfolding end_LE_def by auto lemma end_halt: "\x > 0; inv_end x (Suc 0, l, r)\ \ \ stp. is_final (steps (Suc 0, l, r) (tm_copy_end_new, 0) stp)" proof(rule halt_lemma[OF wf_end_le]) assume great: "0 < x" and inv_start: "inv_end x (Suc 0, l, r)" show "\n. \ is_final (steps (Suc 0, l, r) (tm_copy_end_new, 0) n) \ (steps (Suc 0, l, r) (tm_copy_end_new, 0) (Suc n), steps (Suc 0, l, r) (tm_copy_end_new, 0) n) \ end_LE" proof(rule_tac allI, rule_tac impI) fix n assume notfinal: "\ is_final (steps (Suc 0, l, r) (tm_copy_end_new, 0) n)" obtain s' l' r' where d: "steps (Suc 0, l, r) (tm_copy_end_new, 0) n = (s', l', r')" apply(case_tac "steps (Suc 0, l, r) (tm_copy_end_new, 0) n", auto) done hence "inv_end x (s', l', r') \ s' \ 0" using great inv_start notfinal apply(drule_tac stp = n in inv_end_steps, auto) done hence "(step (s', l', r') (tm_copy_end_new, 0), s', l', r') \ end_LE" apply(cases r'; cases "hd r'") apply(auto simp: inv_end.simps step.simps tm_copy_end_new_def numeral_eqs_upto_12 end_LE_def split: if_splits) done thus "(steps (Suc 0, l, r) (tm_copy_end_new, 0) (Suc n), steps (Suc 0, l, r) (tm_copy_end_new, 0) n) \ end_LE" using d by simp qed qed lemma end_correct: "n > 0 \ \inv_end1 n\ tm_copy_end_new \inv_end0 n\" proof(rule_tac Hoare_haltI) fix l r assume h: "0 < n" "inv_end1 n (l, r)" then have "\ stp. is_final (steps0 (1, l, r) tm_copy_end_new stp)" by (simp add: end_halt inv_end.simps) then obtain stp where "is_final (steps0 (1, l, r) tm_copy_end_new stp)" .. moreover have "inv_end n (steps0 (1, l, r) tm_copy_end_new stp)" apply(rule_tac inv_end_steps) using h by(simp_all add: inv_end.simps) ultimately show "\stp. is_final (steps (1, l, r) (tm_copy_end_new, 0) stp) \ inv_end0 n holds_for steps (1, l, r) (tm_copy_end_new, 0) stp" using h apply(rule_tac x = stp in exI) apply(cases "(steps0 (1, l, r) tm_copy_end_new stp)") apply(simp add: inv_end.simps) done qed (* ---------- tm_weak_copy ---------------------------------------- *) (* This is the original weak variant, with a slightly modified end *) (* --------------------------------------------------------------- *) definition tm_weak_copy :: "instr list" where "tm_weak_copy \ (tm_copy_begin_orig |+| tm_copy_loop_orig) |+| tm_copy_end_new" (* The tm_weak_copy machine and all of its parts are composable * (in the old terminologies: well-formed) *) lemma [intro]: "composable_tm (tm_copy_begin_orig, 0)" "composable_tm (tm_copy_loop_orig, 0)" "composable_tm (tm_copy_end_new, 0)" by (auto simp: composable_tm.simps tm_copy_end_new_def tm_copy_loop_orig_def tm_copy_begin_orig_def) lemma composable_tm0_tm_weak_copy[intro, simp]: "composable_tm0 tm_weak_copy" by (auto simp: tm_weak_copy_def) lemma tm_weak_copy_correct_pre: assumes "0 < x" shows "\inv_begin1 x\ tm_weak_copy \inv_end0 x\" proof - have "\inv_begin1 x\ tm_copy_begin_orig \inv_begin0 x\" by (metis assms begin_correct) moreover have "inv_begin0 x \ inv_loop1 x" unfolding assert_imp_def unfolding inv_begin0.simps inv_loop1.simps unfolding inv_loop1_loop.simps inv_loop1_exit.simps apply(auto simp add: numeral_eqs_upto_12 Cons_eq_append_conv) by (rule_tac x = "Suc 0" in exI, auto) ultimately have "\inv_begin1 x\ tm_copy_begin_orig \inv_loop1 x\" by (rule_tac Hoare_consequence) (auto) moreover have "\inv_loop1 x\ tm_copy_loop_orig \inv_loop0 x\" by (metis assms loop_correct) ultimately have "\inv_begin1 x\ (tm_copy_begin_orig |+| tm_copy_loop_orig) \inv_loop0 x\" by (rule_tac Hoare_plus_halt) (auto) moreover have "\inv_end1 x\ tm_copy_end_new \inv_end0 x\" by (metis assms end_correct) moreover have "inv_loop0 x = inv_end1 x" by(auto simp: inv_end1.simps inv_loop1.simps assert_imp_def) ultimately show "\inv_begin1 x\ tm_weak_copy \inv_end0 x\" unfolding tm_weak_copy_def by (rule_tac Hoare_plus_halt) (auto) qed lemma tm_weak_copy_correct: shows "\\tap. tap = ([]::cell list, Oc \ (Suc n))\ tm_weak_copy \\tap. tap= ([Bk], <(n, n::nat)>)\" proof - have "\inv_begin1 (Suc n)\ tm_weak_copy \inv_end0 (Suc n)\" by (rule tm_weak_copy_correct_pre) (simp) moreover have "(\tap. tap = ([]::cell list, Oc \ (Suc n))) = inv_begin1 (Suc n)" by (auto) moreover have "inv_end0 (Suc n) = (\tap. tap= ([Bk], <(n, n::nat)>))" unfolding fun_eq_iff by (auto simp add: inv_end0.simps tape_of_nat_def tape_of_prod_def) ultimately show "\\tap. tap = ([]::cell list, Oc \ (Suc n))\ tm_weak_copy \\tap. tap= ([Bk], <(n, n::nat)>)\" by simp qed (* additional variants of tm_weak_copy_correct *) lemma tm_weak_copy_correct5: "\\tap. tap = ([], <[n::nat]>)\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, <[n, n]> @ Bk \ l) \" proof - have "\\tap. tap = ([], )\ tm_weak_copy \\tap. tap = ([Bk], <(n, n)>)\" using tape_of_list_def tape_of_nat_def tape_of_nat_list.simps(2) tm_weak_copy_correct by auto then have "\\tap. tap = ([], )\ tm_weak_copy \\tap. tap = ([Bk], <[n, n]>)\" proof - assume "\\tap. tap = ([], )\ tm_weak_copy \\tap. tap = ([Bk], <(n, n)>)\" moreover have "<(n, n)> = <[n, n]>" by (simp add: tape_of_list_def tape_of_nat_list.elims tape_of_prod_def) ultimately show ?thesis by auto qed then have "\\tap. tap = ([], <[n::nat]>)\ tm_weak_copy \\tap. tap = ([Bk], <[n, n]>)\" by (metis tape_of_list_def tape_of_nat_list.simps(2)) then show ?thesis - by (smt Hoare_haltE Hoare_haltI One_nat_def append.right_neutral - holds_for.elims(2) holds_for.simps replicate.simps(1) replicate.simps(2)) + by (smt (verit, del_insts) Hoare_halt_iff append_Nil2 empty_replicate replicate_Suc) qed lemma tm_weak_copy_correct6: "\\tap. \z4. tap = (Bk \ z4, <[n::nat]> @[Bk])\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, <[n::nat, n]> @ Bk \ l) \" proof - have "\\tap. tap = ([], <[n::nat]>)\ tm_weak_copy \\tap. \k l. tap = (Bk \ k, <[n::nat, n]> @ Bk \ l) \" by (rule tm_weak_copy_correct5) then have "\\tap. \kl ll. tap = (Bk \ kl, <[n::nat]> @ Bk \ ll)\ tm_weak_copy \\tap. \kr lr. tap = (Bk \ kr, <[n::nat, n]> @ Bk \ lr)\" using TMC_has_num_res_list_without_initial_Bks_imp_TMC_has_num_res_list_after_adding_Bks_to_initial_left_and_right_tape by auto then have "\\tap. \z4. tap = (Bk \ z4, <[n::nat]> @ Bk \ (Suc 0))\ tm_weak_copy \\tap. \kr lr. tap = (Bk \ kr, <[n::nat, n]> @ Bk \ lr)\" - by (smt Hoare_haltE Hoare_haltI) + by (smt (verit) Hoare_haltE Hoare_haltI) then show ?thesis by auto qed (* ---------------------------------------------------- *) (* Properties of tm_weak_copy for lists of length \ 1 *) (* ---------------------------------------------------- *) (* FABR: Some principle reminder about rewriting/simplification of properties for composed TMs * value "steps0 (1, [Bk,Bk], [Bk]) tm_weak_copy 3" * > "(0, [Bk, Bk, Bk, Bk], [])" * * value "steps0 (1, [Bk,Bk], [Bk]) tm_weak_copy 3 = (0, [Bk, Bk, Bk, Bk], [])" * > True * * 1) If we define a TM directly by an instruction list, the simplifier is able * to compute results (as is the 'value' instruction). * * 2) However, if we define TMs by seq_tm, we must add more rules for rewriting * * Examples: * * Machine tm_strong_copy_post is identical to tm_weak_copy * However, * tm_strong_copy_post is defined as a plain instruction list * tm_weak_copy is defined by seqeuntial composition * * "tm_weak_copy \ (tm_copy_begin_orig |+| tm_copy_loop_orig) |+| tm_copy_end_new" * *) definition strong_copy_post :: "instr list" where "strong_copy_post \ [ (WB,5),(R,2), (R,3),(R,2), (WO,3),(L,4), (L,4),(L,5), (R,11),(R,6), (R,7),(WB,6), (R,7),(R,8), (WO,9),(R,8), (L,10),(L,9), (L,10),(L,5), (R,0),(R,12), (WO,13),(L,14), (R,12),(R,12), (L,15),(WB,14), (R,0),(L,15) ]" value "steps0 (1, [Bk,Bk], [Bk]) strong_copy_post 3 = (0::nat, [Bk, Bk, Bk, Bk], [])" lemma "steps0 (1, [Bk,Bk], [Bk]) strong_copy_post 3 = (0::nat, [Bk, Bk, Bk, Bk], [])" by (simp add: step.simps steps.simps numeral_eqs_upto_12 strong_copy_post_def) (* now, we try the same proof for tm_weak_copy *) (* The rewriter is not able to solve the problem like this: lemma "steps0 (1, [Bk,Bk], [Bk]) tm_weak_copy 3 = (0::nat, [Bk, Bk, Bk, Bk], [])" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_weak_copy_def strong_copy_post_def tm_copy_begin_orig_def tm_copy_loop_orig_def tm_copy_end_new_def adjust.simps shift.simps seq_tm.simps) *) (* There are at least the following two solutions: *) (* Solution 1: * Prove equality "tm_weak_copy = strong_copy_post" * and rewrite with respect to plain instruction list strong_copy_post *) lemma tm_weak_copy_eq_strong_copy_post: "tm_weak_copy = strong_copy_post" unfolding tm_weak_copy_def strong_copy_post_def tm_copy_begin_orig_def tm_copy_loop_orig_def tm_copy_end_new_def by (simp add: adjust.simps shift.simps seq_tm.simps) lemma tm_weak_copy_correct11: "\\tap. tap = ([Bk,Bk], [Bk]) \ tm_weak_copy \\tap. tap = ([Bk,Bk,Bk,Bk],[]) \" proof - have "steps0 (1, [Bk,Bk], [Bk]) strong_copy_post 3 = (0::nat, [Bk, Bk, Bk, Bk], [])" by (simp add: step.simps steps.simps numeral_eqs_upto_12 strong_copy_post_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed lemma tm_weak_copy_correct12: (* this lemma is not used anywhere *) "\\tap. tap = ([Bk,Bk], [Bk]) \ tm_weak_copy \\tap. \k l. tap = ( Bk \ k, Bk \ l) \" proof - have "Bk \ 4 = [Bk, Bk, Bk, Bk]" by (simp add: numeral_eqs_upto_12) moreover have "Bk \ 0 = []" by (simp add: numeral_eqs_upto_12) ultimately have "\\tap. tap = ([Bk,Bk], [Bk]) \ tm_weak_copy \\tap. tap = ( Bk \ 4, Bk \ 0) \" using tm_weak_copy_correct11 by auto then show ?thesis by (metis (no_types, lifting) Hoare_halt_def holds_for.elims(2) holds_for.simps) qed lemma tm_weak_copy_correct13: "\\tap. tap = ([], [Bk,Bk]@r) \ tm_weak_copy \\tap. tap = ([Bk,Bk], r) \" proof - have "steps0 (1, [], [Bk,Bk]@r) strong_copy_post 3 = (0::nat, [Bk,Bk], r)" by (simp add: step.simps steps.simps numeral_eqs_upto_12 strong_copy_post_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed (* Solution 2: (pays off for more complicated cases) * * Use compositional reasoning on components of tm_weak_copy using our Hoare rules." * * "tm_weak_copy \ (tm_copy_begin_orig |+| tm_copy_loop_orig) |+| tm_copy_end_new" * * See our exploration in file HaskellCode/EngineeringOf_StrongCopy.hs * These yield the concrete steps and results to compute for the intermediate * steps in the following proofs. * *) (* A compositional proof using Hoare_rules for * * "\\tap. tap = ([Bk,Bk], [Bk]) \ tm_weak_copy \\tap. tap = ([Bk,Bk,Bk,Bk],[]) \" * * "steps0 (1, [Bk, Bk], [Bk]) tm_copy_begin_orig 1 = (0, [Bk, Bk], [Bk])" * "steps0 (1, [Bk, Bk], [Bk]) tm_copy_loop_orig 1 = (0, [Bk, Bk, Bk], [] )" * "steps0 (1, [Bk, Bk, Bk], [] ) tm_copy_end_new 1 = (0, [Bk, Bk, Bk, Bk], [] )" * * See HaskellCode/EngineeringOf_StrongCopy.hs for the exploration. *) lemma tm_weak_copy_correct11': "\\tap. tap = ([Bk,Bk], [Bk]) \ tm_weak_copy \\tap. tap = ([Bk,Bk,Bk,Bk],[]) \" proof - have "\\tap. tap = ([Bk,Bk], [Bk]) \ (tm_copy_begin_orig |+| tm_copy_loop_orig) |+| tm_copy_end_new \\tap. tap = ([Bk,Bk,Bk,Bk],[]) \" proof (rule Hoare_plus_halt) show "composable_tm0 (tm_copy_begin_orig |+| tm_copy_loop_orig)" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_copy_begin_orig_def tm_copy_loop_orig_def) next show "\\tap. tap = ([Bk, Bk, Bk], [])\ tm_copy_end_new \\tap. tap = ([Bk, Bk, Bk, Bk], [])\" proof - have "steps0 (1, [Bk, Bk, Bk], []) tm_copy_end_new 1 = (0, [Bk, Bk, Bk, Bk], [])" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_copy_end_new_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed next show " \\tap. tap = ([Bk, Bk], [Bk])\ tm_copy_begin_orig |+| tm_copy_loop_orig \\tap. tap = ([Bk, Bk, Bk], [])\" proof (rule Hoare_plus_halt) show "composable_tm0 tm_copy_begin_orig" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_copy_begin_orig_def) next show "\\tap. tap = ([Bk, Bk], [Bk])\ tm_copy_begin_orig \\tap. tap = ([Bk, Bk], [Bk])\" proof - have "steps0 (1, [Bk, Bk], [Bk]) tm_copy_begin_orig 1 = (0, [Bk, Bk], [Bk])" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_copy_begin_orig_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed next show "\\tap. tap = ([Bk, Bk], [Bk])\ tm_copy_loop_orig \\tap. tap = ([Bk, Bk, Bk], [])\" proof - have "steps0 (1, [Bk, Bk], [Bk]) tm_copy_loop_orig 1 = (0, [Bk, Bk, Bk], [])" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_copy_loop_orig_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed qed qed then show ?thesis unfolding tm_weak_copy_def by auto qed (* -- *) (* A compositional proof using Hoare_rules for * * "\\tap. tap = ([], [Bk,Bk]@r) \ tm_weak_copy \\tap. \k l. tap = ([Bk,Bk], r) \" * * "steps0 (1, [], [Bk,Bk] @ r) tm_copy_begin_orig 1 = (0, [], [Bk,Bk] @ r)" * "steps0 (1, [], [Bk,Bk] @ r) tm_copy_loop_orig 1 = (0, [Bk], [Bk] @ r)" * "steps0 (1, [Bk], [Bk] @ r) tm_copy_end_new 1 = (0, [Bk, Bk], r)" * * See HaskellCode/EngineeringOf_StrongCopy.hs for the exploration. *) lemma tm_weak_copy_correct13': "\\tap. tap = ([], [Bk,Bk]@r) \ tm_weak_copy \\tap. tap = ([Bk,Bk], r) \" proof - have "\\tap. tap = ([], [Bk,Bk]@r) \ (tm_copy_begin_orig |+| tm_copy_loop_orig) |+| tm_copy_end_new \\tap. tap = ([Bk,Bk], r) \" proof (rule Hoare_plus_halt) show "composable_tm0 (tm_copy_begin_orig |+| tm_copy_loop_orig)" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_copy_begin_orig_def tm_copy_loop_orig_def) next show "\\tap. tap = ([Bk], [Bk] @ r)\ tm_copy_end_new \\tap. tap = ([Bk, Bk], r)\" proof - have "steps0 (1, [Bk], [Bk] @ r) tm_copy_end_new 1 = (0, [Bk, Bk], r)" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_copy_end_new_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed next show " \\tap. tap = ([], [Bk, Bk] @ r)\ tm_copy_begin_orig |+| tm_copy_loop_orig \\tap. tap = ([Bk], [Bk] @ r)\" proof (rule Hoare_plus_halt) show "composable_tm0 tm_copy_begin_orig" by (simp add: composable_tm.simps adjust.simps shift.simps seq_tm.simps tm_copy_begin_orig_def) next show "\\tap. tap = ([], [Bk, Bk] @ r)\ tm_copy_begin_orig \\tap. tap = ([], [Bk,Bk] @ r)\" proof - have "steps0 (1, [], [Bk, Bk] @ r) tm_copy_begin_orig 1 = (0, [], [Bk,Bk] @ r)" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_copy_begin_orig_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed next show "\\tap. tap = ([], [Bk, Bk] @ r)\ tm_copy_loop_orig \\tap. tap = ([Bk], [Bk] @ r)\" proof - have "steps0 (1, [], [Bk,Bk] @ r) tm_copy_loop_orig 1 = (0, [Bk], [Bk] @ r)" by (simp add: step.simps steps.simps numeral_eqs_upto_12 tm_copy_loop_orig_def) then show ?thesis - by (smt Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) + by (smt (verit) Hoare_haltI holds_for.simps is_final_eq tm_weak_copy_eq_strong_copy_post) qed qed qed then show ?thesis unfolding tm_weak_copy_def by auto qed end