Theory StrongCopyTM
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›.
›
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) ]"
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 (verit) Hoare_haltI holds_for.simps is_final_eq)
qed
corollary tm_skip_first_arg_correct_Nil':
"length nl = 0
⟹ ⦃λtap. tap = ([], <nl::nat list> )⦄ tm_skip_first_arg ⦃λtap. tap = ([], [Bk]) ⦄"
using tm_skip_first_arg_correct_Nil
by (simp add: tm_skip_first_arg_correct_Nil )
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
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)
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
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
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
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 (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
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)
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"
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
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
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 = ([], <nl::nat list> )⦄ 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
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] @ (<ns::nat list>) )"
| "inv_tm_skip_first_arg_len_gt_1_s2 n ns (l, r) =
(∃n1 n2. l = Oc ↑ (Suc n1) ∧ r = Oc ↑ n2 @ [Bk] @ (<ns::nat list>) ∧
Suc n1 + n2 = Suc n)"
| "inv_tm_skip_first_arg_len_gt_1_s3 n ns (l, r) = (
l = Bk # Oc ↑ (Suc n) ∧ r = (<ns::nat list>)
)"
| "inv_tm_skip_first_arg_len_gt_1_s0 n ns (l, r) = (
l = Bk# Oc ↑ (Suc n) ∧ r = (<ns::nat list>)
)"
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] @ (<ns::nat list>)"
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] @ (<ns::nat list>) ∧ Suc n1 + n2 = Suc n)"
by auto
then obtain n1 n2
where w_n1_n2: "l = Oc ↑ (Suc n1) ∧ r = Oc ↑ n2 @ [Bk] @ (<ns::nat list>) ∧ 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] @ (<ns::nat list>) ∧ Suc n1 = Suc n"
by auto
then have "step0 (2, Oc ↑ (Suc n1), [Bk] @ (<ns::nat list>)) tm_skip_first_arg = (3, Bk # Oc ↑ (Suc n1), (<ns::nat list>))"
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), (<ns::nat list>))"
proof -
from ‹l = Oc ↑ (Suc n1) ∧ r = [Bk] @ (<ns::nat list>) ∧ 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), (<ns::nat list>))" 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] @ (<ns::nat list>) ∧ Suc n1 + n2 = Suc n"
by auto
then have "step0 (2, Oc ↑ (Suc n1), Oc ↑ Suc n2' @ [Bk] @ (<ns::nat list>)) tm_skip_first_arg
= (2, Oc # Oc ↑ (Suc n1), Oc ↑ n2' @ [Bk] @ (<ns::nat list>))"
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] @ (<ns::nat list>))"
proof -
from ‹l = Oc ↑ (Suc n1) ∧ r = Oc ↑ Suc n2' @ [Bk] @ (<ns::nat list>) ∧ Suc n1 + n2 = Suc n›
have "Oc # Oc ↑ (Suc n1) = Oc ↑ Suc (Suc n1) ∧ Oc ↑ n2' @ [Bk] @ (<ns::nat list>)
= Oc ↑ n2' @ [Bk] @ (<ns::nat list>) ∧ 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] @ (<ns::nat list>))"
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] @ (<ns::nat list>) ∧ 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 = (<ns::nat list>)"
by auto
moreover with ‹length ns > 0› have "(ns::nat list) ≠ [] ∧ hd (<ns::nat list>) = Oc"
using numeral_list_head_is_Oc
by force
moreover from this have "<ns::nat list> = Oc#(tl (<ns::nat list>))"
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), (<ns::nat list>)) tm_skip_first_arg
= (0, Bk # Oc ↑ (Suc n), (<ns::nat list>))"
proof -
from ‹<ns> = Oc # tl (<ns>)›
have "step0 (3, Bk # Oc ↑ (Suc n), (<ns::nat list>)) tm_skip_first_arg
= step0 (3, Bk # Oc ↑ (Suc n), Oc # tl (<ns>)) tm_skip_first_arg"
by auto
also have "... = (0, Bk # Oc ↑ (Suc n), Oc # tl (<ns>))"
by (simp add: tm_skip_first_arg_def step.simps steps.simps numeral_eqs_upto_12)
also with ‹<ns> = Oc # tl (<ns>)› have "... = (0, Bk # Oc ↑ (Suc n), (<ns::nat list>))" by auto
finally show "step0 (3, Bk # Oc ↑ (Suc n), (<ns::nat list>)) tm_skip_first_arg
= (0, Bk # Oc ↑ (Suc n), (<ns::nat list>))"
by auto
qed
moreover with ‹l = Bk # Oc ↑ (Suc n) ∧ r = (<ns::nat list>)›
have "inv_tm_skip_first_arg_len_gt_1 n ns (0, Bk # Oc ↑ (Suc n), (<ns::nat list>))"
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] @ (<ns::nat list>) ) tm_skip_first_arg stp)"
and "0 < length ns"
shows "⦃λtap. tap = ([], Oc ↑ Suc n @ [Bk] @ (<ns::nat list>) ) ⦄
tm_skip_first_arg
⦃ λtap. tap = (Bk# Oc ↑ Suc n, (<ns::nat list>) ) ⦄"
proof (rule Hoare_consequence)
show " (λtap. tap = ([], Oc ↑ Suc n @ [Bk] @ (<ns::nat list>)))
↦ (λtap. tap = ([], Oc ↑ Suc n @ [Bk] @ (<ns::nat list>)))"
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, <ns>))"
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] @ <ns>)⦄ 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] @ <ns::nat list>)"
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 (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
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] @ <ns::nat list>) tm_skip_first_arg stp)"
proof -
assume A: "0 < length ns"
show "∃stp. is_final (steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns::nat list>) 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] @ <ns>) tm_skip_first_arg stp)" .
have INV: "inv_tm_skip_first_arg_len_gt_1 n ns (steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns>) 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] @ <ns>)"
by (simp add: tape_of_list_def tape_of_nat_def)
qed
have SUC_STEP_RED:
"steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns>) tm_skip_first_arg (Suc stp) =
step0 (steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns>) tm_skip_first_arg stp) tm_skip_first_arg"
by (rule step_red)
show "( steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns>) tm_skip_first_arg (Suc stp),
steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns>) tm_skip_first_arg stp
) ∈ measure_tm_skip_first_arg_len_gt_1"
proof (cases "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns>) tm_skip_first_arg stp")
case (fields s l r2)
then have
cf_cases: "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ <ns>) 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"
with cf_cases not_final
show ?thesis by auto
next
assume "s=4"
with cf_cases not_final INV
show ?thesis by auto
next
assume "s=5"
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] @ (<ns>)) 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] @ (<ns>)"
by auto
with cf_cases and ‹s=1› and SUC_STEP_RED
have "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg (Suc stp) =
step0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg"
by auto
also have "... = (2,[Oc],Oc ↑ n @ [Bk] @ (<ns>))"
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] @ (<ns>)) tm_skip_first_arg (Suc stp)
= (2,[Oc],Oc ↑ n @ [Bk] @ (<ns>))"
by auto
with ‹steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) 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] @ (<ns>)) 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] @ (<ns::nat list>) ∧
Suc n1 + n2 = Suc n)"
by auto
then obtain n1 n2 where
w_n1_n2: "l = Oc ↑ (Suc n1) ∧ r2 = Oc ↑ n2 @ [Bk] @ (<ns::nat list>) ∧ 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] @ (<ns::nat list>)" by auto
with ‹steps0 (1, [],Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (2, l, r2)›
have "steps0 (1, [],Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (2, l, [Bk] @ (<ns::nat list>))"
by auto
with SUC_STEP_RED
have "steps0 (1, [],Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg (Suc stp) =
step0 (2, l, [Bk] @ (<ns::nat list>)) tm_skip_first_arg"
by auto
also have "... = (3,Bk#l,(<ns>))"
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] @ (<ns>)) tm_skip_first_arg (Suc stp) = (3,Bk#l,(<ns>))"
by auto
with ‹steps0 (1, [],Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (2, l, [Bk] @ (<ns::nat list>))›
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#(<ns>)" by auto
with ‹steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (2, l, r2)›
have "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (2, l, Oc ↑ Suc n2'@Bk#(<ns>))"
by auto
with SUC_STEP_RED
have "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg (Suc stp) =
step0 (2, l, Oc ↑ Suc n2'@Bk#(<ns>)) tm_skip_first_arg"
by auto
also have "... = (2,Oc#l,Oc ↑ n2'@Bk#(<ns>))"
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] @ (<ns>)) tm_skip_first_arg (Suc stp)
= (2,Oc#l,Oc ↑ n2'@Bk#(<ns>))"
by auto
with ‹steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (2, l, Oc ↑ Suc n2'@Bk#(<ns>))›
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] @ (<ns>)) 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 = (<ns::nat list>)"
by auto
with ‹steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (3, l, r2)›
have "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (3, Bk # Oc ↑ (Suc n), (<ns::nat list>))"
by auto
with SUC_STEP_RED
have "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg (Suc stp) =
step0 (3, Bk # Oc ↑ (Suc n), (<ns::nat list>)) tm_skip_first_arg"
by auto
also have "... = (0, Bk # Oc ↑ (Suc n),(<ns::nat list>))"
proof -
from ‹length ns > 0› have "(ns::nat list) ≠ [] ∧ hd (<ns::nat list>) = Oc"
using numeral_list_head_is_Oc
by force
then have "<ns::nat list> = Oc#(tl (<ns::nat list>))"
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), (<ns::nat list>)) tm_skip_first_arg
= step0 (3, Bk # Oc ↑ (Suc n), Oc#(tl (<ns::nat list>))) tm_skip_first_arg"
by auto
also have "... = (0, Bk # Oc ↑ (Suc n), Oc#(tl (<ns::nat list>)))"
by (auto simp add: tm_skip_first_arg_def numeral_eqs_upto_12 step.simps steps.simps)
also with ‹<ns::nat list> = Oc#(tl (<ns::nat list>))›
have "... = (0, Bk # Oc ↑ (Suc n), <ns::nat list>)"
by auto
finally show "step0 (3, Bk # Oc ↑ (Suc n), (<ns::nat list>)) tm_skip_first_arg
= (0, Bk # Oc ↑ (Suc n), <ns::nat list>)" by auto
qed
finally have "steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg (Suc stp) =
(0, Bk # Oc ↑ (Suc n),(<ns::nat list>))"
by auto
with ‹steps0 (1, [], Oc ↑ Suc n @ [Bk] @ (<ns>)) tm_skip_first_arg stp = (3, Bk # Oc ↑ (Suc n), (<ns::nat list>))›
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] @ (<ns::nat list>) ) ⦄
tm_skip_first_arg
⦃ λtap. tap = (Bk# Oc ↑ Suc n, (<ns::nat list>) ) ⦄"
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] @ (<ns::nat list>) ) 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 = ([], <nl::nat list> )⦄ tm_skip_first_arg ⦃ λtap. tap = (Bk# <rev [hd nl]>, <tl nl>) ⦄"
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 "<nl::nat list> = <hd nl # tl nl>"
by auto
also with ‹tl nl ≠ []› have "... = <hd nl> @ [Bk] @ (<tl nl>)"
by (simp add: tape_of_nat_list_cons_eq)
also with ‹(nl::nat list) ≠ []› have "... = Oc ↑ Suc (hd nl) @ [Bk] @ (<tl nl>)"
using tape_of_nat_def by blast
finally have "<nl::nat list> = Oc ↑ Suc (hd nl) @ [Bk] @ (<tl nl>)"
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] @ (<tl nl>) ) ⦄
tm_skip_first_arg
⦃ λtap. tap = (Bk# Oc ↑ Suc (hd nl), (<tl nl>) ) ⦄"
using tm_skip_first_arg_len_gt_1_total_correctness_pre
by force
with ‹<nl::nat list> = Oc ↑ Suc (hd nl) @ [Bk] @ (<tl nl>)› have
"⦃λtap. tap = ([], <nl::nat list> ) ⦄
tm_skip_first_arg
⦃ λtap. tap = (Bk# Oc ↑ Suc (hd nl), (<tl nl>) ) ⦄"
by force
moreover have "<rev [hd nl]> = 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
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)
]"
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 (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
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"
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
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)) ) )"
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 ≠ [])
)"
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
)"
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
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)
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 -
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
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
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
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
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
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
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
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
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
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]"
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"
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"
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'"
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
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]"
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
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
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]"
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"
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
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, [] , 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
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]"
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
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"
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"
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
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
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"
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)"
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)"
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"
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"
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"
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)"
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)
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 = []"
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)
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)"
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)
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"
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"
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)
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' ≠ []"
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"
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)
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' ≠ []"
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
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
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
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
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 (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 (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 (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
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"
with cf_at_stp not_final
show ?thesis by auto
next
assume "s=1"
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
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
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
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
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"
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
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
with cf_at_current show ?thesis
by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def)
next
assume "s=3"
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
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
with cf_at_current show ?thesis
by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def)
next
assume "s=5"
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
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
with cf_at_current show ?thesis
by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def)
next
assume "s=6"
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
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
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
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"
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
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
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
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
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"
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
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
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"
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
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
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
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
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 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
with cf_at_current show ?thesis
by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def)
qed
next
assume "s=10"
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" .
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
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" .
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
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
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
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)
case Nil
then have minor: "CL = []" .
with unpacked_INV cf_at_stp and ‹s=10› and SUC_STEP_RED
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))"
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
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
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)"
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
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
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)"
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
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"
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]"
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
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]"
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
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]"
then obtain rex where
unpacked_INV: "l = [Bk] ∧ r = rev [Oc] @ Oc # Bk ↑ Suc rex ∧ CL = [Oc, Bk]" by blast
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
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 ≠ []"
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]"
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
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 = []"
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
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"
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
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
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 = []"
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
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]"
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
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"
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
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"
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
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]"
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
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
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)"
then obtain rex where
unpacked_INV: "l = [] ∧ r = Bk # rev CL @ Oc # Bk ↑ Suc rex ∧ (CL = [] ∨ last CL = Oc)" by blast
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
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"
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)
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
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"
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)
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
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"
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)
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 = []"
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
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]"
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
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"
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
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"
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
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
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
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
with cf_at_current show ?thesis
by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def)
next
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
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
with cf_at_current show ?thesis
by (auto simp add: measure_tm_erase_right_then_dblBk_left_erp_def)
qed
qed
qed
qed
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
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(<take n nl>) ∧ noDblBk CL ∧ CL = [] ∧
CR = (<drop n nl>) ∧ noDblBk CR"
proof -
have "rev(<take n nl>) = <rev(take n nl)>"
by (rule rev_numeral_list)
also with assms have "... = <rev( butlast (take n nl) @ [last (take n nl)] )>"
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(<take n nl>) = <(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(<take n nl>) = [Oc]" by auto
then have "[Oc] @ [] = rev(<take n nl>) ∧ noDblBk [] ∧ ([] = [] ∨ [] ≠ [] ∧ last [] = Oc) ∧
(<drop n nl>) = (<drop n nl>) ∧ noDblBk (<drop n nl>)"
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(<take n nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<drop n nl>) ∧ noDblBk CR"
proof -
have minor: "rev(<take n nl>) = <rev(take n nl)>"
by (rule rev_numeral_list)
also with assms have "... = <rev( butlast (take n nl) @ [last (take n nl)] )>"
by simp
also have "... = <(rev[last (take n nl)]) @ (rev (butlast (take n nl)))>"
by simp
finally have major: "rev(<take n nl>) = <(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(<take n nl>) = <(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(<take n nl>) = 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(<take n nl>) ∧
noDblBk (Oc↑ (last (take n nl))) ∧ (Oc↑ (last (take n nl))) ≠ [] ∧ last (Oc↑ (last (take n nl))) = Oc ∧
(<drop n nl>) = (<drop n nl>) ∧ noDblBk (<drop n nl>)" using assms
by (simp add: noDblBk_Bk_Oc_rep noDblBk_tape_of_nat_list)
then show ?thesis by auto
qed
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(<hd nl>) ∧ noDblBk CL ∧ CL = [] ∧
CR = (<tl nl>) ∧ 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(<take n nl>) ∧ noDblBk CL ∧ CL = [] ∧
CR = (<drop n nl>) ∧ 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 (<take n nl>) ∧ noDblBk CL ∧ CL = [] ∧ CR = <drop n nl> ∧ 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(<hd nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<tl nl>) ∧ 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(<take n nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<drop n nl>) ∧ 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
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<n"
and "n ≤ length nl"
and "last (take n nl) = 0"
shows "∃CL CR.
[Oc] @ CL = rev(<take n nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<drop n nl>) ∧ noDblBk CR"
proof -
have minor: "rev(<take n nl>) = <rev(take n nl)>"
by (rule rev_numeral_list)
also with assms have "... = <rev( butlast (take n nl) @ [last (take n nl)] )>"
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(<take n nl>) = <(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(<take n nl>) = [Oc] @ [Bk] @ <rev (butlast (take n nl))>"
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(<take n nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<drop n nl>) ∧ noDblBk CR"
proof -
have minor: "rev(<take n nl>) = <rev(take n nl)>"
by (rule rev_numeral_list)
also with assms have "... = <rev( butlast (take n nl) @ [last (take n nl)] )>"
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(<take n nl>) = <(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(<take n nl>) = <(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(<take n nl>) = 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(<take n nl>) ∧ 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 ∧
(<drop n nl>) = (<drop n nl>) ∧ noDblBk (<drop n nl>)"
using assms
‹<rev [last (take n nl)]> = Oc ↑ Suc (last (take n nl))›
‹Oc ↑ Suc (last (take n nl)) = Oc # Oc ↑ last (take n nl)›
‹butlast (take n nl) ≠ []› ‹rev (<take n nl>) = <rev [last (take n nl)]> @ [Bk] @ <rev (butlast (take n nl))>›
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<n"
and "n ≤ length nl"
shows "∃CL CR.
[Oc] @ CL = rev(<take n nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<drop n nl>) ∧ 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
lemma tm_erase_right_then_dblBk_left_erp_total_correctness_one_arg:
assumes "1 ≤ length (nl::nat list)"
shows "⦃ λtap. tap = (Bk# rev(<hd nl>), <tl nl>) ⦄
tm_erase_right_then_dblBk_left
⦃ λtap. ∃rex. tap = ([], [Bk,Bk] @ (<hd nl>) @ [Bk] @ Bk ↑ rex ) ⦄"
proof (cases "hd nl")
case 0
then have "hd nl = 0" .
with assms
have "∃CL CR.
[Oc] @ CL = rev(<hd nl>) ∧ noDblBk CL ∧ CL = [] ∧
CR = (<tl nl>) ∧ 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(<hd nl>) ∧ noDblBk CL ∧ CL = [] ∧
CR = (<tl nl>) ∧ noDblBk CR" by blast
show ?thesis
proof (rule Hoare_consequence)
from assms and w_CL_CR show "(λtap. tap = (Bk # rev (<hd nl>), <tl nl>)) ↦ (λtap. tap = ([Bk,Oc] @ CL, CR))"
using Cons_eq_appendI append_self_conv assert_imp_def by auto
next
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
show "(λtap. ∃rex. tap = ([], [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk ↑ rex)) ↦ (λtap. ∃rex. tap = ([], [Bk, Bk] @ <hd nl> @ [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(<hd nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<tl nl>) ∧ 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(<hd nl>) ∧ noDblBk CL ∧ CL ≠ [] ∧ last CL = Oc ∧
CR = (<tl nl>) ∧ noDblBk CR" by blast
show ?thesis
proof (rule Hoare_consequence)
from assms and w_CL_CR show "(λtap. tap = (Bk # rev (<hd nl>), <tl nl>)) ↦ (λtap. tap = ([Bk,Oc] @ CL, CR))"
by (simp add: w_CL_CR assert_imp_def)
next
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
show "(λtap. ∃rex. tap = ([], [Bk, Bk] @ rev CL @ [Oc, Bk] @ Bk ↑ rex)) ↦ (λtap. ∃rex. tap = ([], [Bk, Bk] @ <hd nl> @ [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
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 = ([], <nl::nat list>) ⦄ tm_check_for_one_arg ⦃λtap. tap = ([Bk,Bk], [Bk] ) ⦄"
proof -
assume major: "length nl = 0"
have "⦃λtap. tap = ([], <nl::nat list>) ⦄ (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 = ([], <nl::nat list>) ⦄ 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 = ([], <nl::nat list>) ⦄ tm_check_for_one_arg ⦃λtap. ∃z4. tap = (Bk ↑ z4, <nl> @ [Bk])⦄"
proof -
assume major: "length nl = 1"
have " ⦃λtap. tap = ([], <nl::nat list>) ⦄
(tm_skip_first_arg |+| tm_erase_right_then_dblBk_left)
⦃λtap. ∃z4. tap = (Bk ↑ z4, <nl> @ [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 = ([], <nl::nat list>) ⦄ 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 = ([], <nl::nat list>) ⦄ tm_skip_first_arg ⦃ λtap. tap = ([Bk], <nl> @[Bk])⦄" using major
by auto
next
from major
have "⦃λtap. tap = ([], <nl> @ [Bk])⦄ tm_erase_right_then_dblBk_left ⦃λtap. tap = ([Bk, Bk], <nl> @ [Bk])⦄"
using tm_erase_right_then_dblBk_left_dnp_total_correctness
by simp
with major have "∃stp. is_final (steps0 (1, [],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp) ∧
(steps0 (1, [],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, [Bk, Bk], <nl> @ [Bk]))"
unfolding Hoare_halt_def
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, [],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp) ∧
(steps0 (1, [],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, [Bk, Bk], <nl> @ [Bk]))" by blast
then have "is_final (steps0 (1, Bk ↑ 0,<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp) ∧
(steps0 (1, Bk ↑ 0,<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk ↑ 2, <nl> @ [Bk]))"
by (simp add: is_finalI numeral_eqs_upto_12(1))
then have "∃z3. z3 ≤ 0 + 1 ∧
is_final (steps0 (1, Bk ↑ (0+1),<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp) ∧
(steps0 (1, Bk ↑ (0+1),<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk ↑ (2+z3), <nl> @ [Bk]))"
by (metis is_finalI steps_left_tape_EnlargeBkCtx)
then have "is_final (steps0 (1, [Bk],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp) ∧
(∃z4. steps0 (1, [Bk],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left stp = (0, Bk ↑ z4, <nl> @ [Bk]))"
by (metis One_nat_def add.left_neutral replicate_0 replicate_Suc)
then have "∃n. is_final (steps0 (1, [Bk],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left n) ∧
(∃z4. steps0 (1, [Bk],<nl::nat list>@ [Bk]) tm_erase_right_then_dblBk_left n = (0, Bk ↑ z4, <nl> @ [Bk]))"
by blast
then show "⦃λtap. tap = ([Bk], <nl::nat list> @ [Bk])⦄
tm_erase_right_then_dblBk_left
⦃λtap. ∃z4. tap = (Bk ↑ z4, <nl::nat list> @ [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 = ([], <nl::nat list> )⦄ tm_check_for_one_arg ⦃ λtap. ∃ l. tap = ( [], [Bk,Bk] @ <[hd nl]> @ Bk ↑ l) ⦄"
proof -
assume major: "length nl > 1"
have " ⦃λtap. tap = ([], <nl::nat list> )⦄
(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 = ([], <nl::nat list> )⦄ tm_skip_first_arg ⦃ λtap. tap = (Bk# <rev [hd nl]>, <tl nl>) ⦄"
using tm_skip_first_arg_len_gt_1_total_correctness
by simp
next
from major
have "⦃ λtap. tap = (Bk# rev(<hd nl>), <tl nl>) ⦄ tm_erase_right_then_dblBk_left ⦃ λtap. ∃rex. tap = ([], [Bk,Bk] @ (<hd nl>) @ [Bk] @ Bk ↑ rex ) ⦄"
using tm_erase_right_then_dblBk_left_erp_total_correctness_one_arg
by simp
then have "⦃ λtap. tap = (Bk# <rev [hd nl]>, <tl nl>) ⦄ 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# <rev [hd nl]>, <tl nl>) ⦄ tm_erase_right_then_dblBk_left ⦃ λtap. ∃rex. tap = ([], [Bk,Bk] @ <[hd nl]> @ Bk ↑ (Suc rex) ) ⦄"
by force
then show "⦃λtap. tap = (Bk # <rev [hd nl]>, <tl nl>)⦄ tm_erase_right_then_dblBk_left ⦃λtap. ∃l. tap = ([], [Bk, Bk] @ <[hd nl]> @ Bk ↑ l)⦄"
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
definition
tm_strong_copy :: "instr list"
where
"tm_strong_copy ≡ tm_check_for_one_arg |+| tm_weak_copy"
lemma tm_strong_copy_total_correctness_Nil:
"length nl = 0
⟹ ⦃λtap. tap = ([], <nl::nat list>) ⦄ tm_strong_copy ⦃λtap. tap = ([Bk,Bk,Bk,Bk],[]) ⦄"
proof -
assume major: "length nl = 0"
have " ⦃λtap. tap = ([], <nl::nat list>) ⦄
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 = ([], <nl::nat list>) ⦄ 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 = ([], <nl::nat list>) ⦄ tm_strong_copy ⦃ λtap. ∃l. tap = ([Bk,Bk], <[hd nl]> @ Bk ↑ l) ⦄"
proof -
assume major: "1 < length nl"
have " ⦃λtap. tap = ([], <nl::nat list>) ⦄
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 = ([], <nl::nat list>) ⦄ 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 (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 = ([], <nl::nat list>) ⦄ tm_strong_copy ⦃ λtap. ∃k l. tap = (Bk ↑ k, <[hd nl, hd nl]> @ Bk ↑ l) ⦄"
proof -
assume major: "1 = length nl"
have " ⦃λtap. tap = ([], <nl::nat list>) ⦄
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 = ([], <nl::nat list>) ⦄ tm_check_for_one_arg ⦃λtap. ∃z4. tap = (Bk ↑ z4, <nl> @ [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 "<nl> = <[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, <nl> @ [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