(* Title: Turing machines
Author: Xu Jian <xujian817@hotmail.com>
Maintainer: Xu Jian
*)
theory turing_basic
imports Main
begin
section {* Basic definitions of Turing machine *}
definition
"iseven n \<equiv> \<exists>x. n = 2 * x"
datatype action = W0 | W1 | L | R | Nop
datatype cell = Bk | Oc
type_synonym tape = "cell list \<times> cell list"
type_synonym state = nat
type_synonym instr = "action \<times> state"
type_synonym tprog = "instr list"
type_synonym config = "state \<times> tape"
fun nth_of where
"nth_of [] _ = None"
| "nth_of (x # xs) 0 = Some x"
| "nth_of (x # xs) (Suc n) = nth_of xs n"
lemma nth_of_map [simp]:
shows "nth_of (map f p) n = (case (nth_of p n) of None \<Rightarrow> None | Some x \<Rightarrow> Some (f x))"
apply(induct p arbitrary: n)
apply(auto)
apply(case_tac n)
apply(auto)
done
fun
fetch :: "tprog \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr"
where
"fetch p 0 b = (Nop, 0)"
| "fetch p (Suc s) Bk =
(case nth_of p (2 * s) of
Some i \<Rightarrow> i
| None \<Rightarrow> (Nop, 0))"
|"fetch p (Suc s) Oc =
(case nth_of p ((2 * s) + 1) of
Some i \<Rightarrow> i
| None \<Rightarrow> (Nop, 0))"
lemma fetch_Nil [simp]:
shows "fetch [] s b = (Nop, 0)"
apply(case_tac s)
apply(auto)
apply(case_tac b)
apply(auto)
done
fun
update :: "action \<Rightarrow> tape \<Rightarrow> tape"
where
"update W0 (l, r) = (l, Bk # (tl r))"
| "update W1 (l, r) = (l, Oc # (tl r))"
| "update L (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))"
| "update R (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))"
| "update Nop (l, r) = (l, r)"
abbreviation
"read r == if (r = []) then Bk else hd r"
fun
step :: "config \<Rightarrow> tprog \<Rightarrow> config"
where
"step (s, l, r) p = (let (a, s') = fetch p s (read r) in (s', update a (l, r)))"
fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config"
where
"steps c p 0 = c" |
"steps c p (Suc n) = steps (step c p) p n"
lemma step_red [simp]:
shows "steps c p (Suc n) = step (steps c p n) p"
by (induct n arbitrary: c) (auto)
lemma steps_add [simp]:
shows "steps c p (m + n) = steps (steps c p m) p n"
by (induct m arbitrary: c) (auto)
definition
tm_wf :: "tprog \<Rightarrow> bool"
where
"tm_wf p = (length p \<ge> 2 \<and> iseven (length p) \<and> (\<forall>(a, s) \<in> set p. s \<le> length p div 2))"
(* FIXME: needed? *)
lemma halt_lemma:
"\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
by (metis wf_iff_no_infinite_down_chain)
abbreviation exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_ \<up> _" [100, 99] 100)
where "x \<up> n == replicate n x"
definition tinres :: "cell list \<Rightarrow> cell list \<Rightarrow> bool"
where
"tinres xs ys = (\<exists>n. xs = ys @ Bk \<up> n \<or> ys = xs @ Bk \<up> n)"
fun
shift :: "tprog \<Rightarrow> nat \<Rightarrow> tprog"
where
"shift p n = (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)"
lemma length_shift [simp]:
"length (shift p n) = length p"
by (simp)
fun
adjust :: "tprog \<Rightarrow> tprog"
where
"adjust p = map (\<lambda> (a, s). (a, if s = 0 then ((length p) div 2) + 1 else s)) p"
lemma length_adjust[simp]:
shows "length (adjust p) = length p"
by (induct p) (auto)
definition
tm_comp :: "tprog \<Rightarrow> tprog \<Rightarrow> tprog" ("_ |+| _" [0, 0] 100)
where
"tm_comp p1 p2 = ((adjust p1) @ (shift p2 ((length p1) div 2)))"
fun
is_final :: "config \<Rightarrow> bool"
where
"is_final (s, l, r) = (s = 0)"
lemma is_final_steps:
assumes "is_final (s, l, r)"
shows "is_final (steps (s, l, r) p n)"
using assms by (induct n) (auto)
fun
holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100)
where
"P holds_for (s, l, r) = P (l, r)"
lemma step_0 [simp]:
shows "step (0, (l, r)) p = (0, (l, r))"
by simp
lemma steps_0 [simp]:
shows "steps (0, (l, r)) p n = (0, (l, r))"
by (induct n) (simp_all)
lemma is_final_holds[simp]:
assumes "is_final c"
shows "Q holds_for (steps c p n) = Q holds_for c"
using assms
apply(induct n)
apply(case_tac [!] c)
apply(auto)
done
type_synonym assert = "tape \<Rightarrow> bool"
definition assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100)
where
"assert_imp P Q = (\<forall>l r. P (l, r) \<longrightarrow> Q (l, r))"
lemma holds_for_imp:
assumes "P holds_for c"
and "P \<mapsto> Q"
shows "Q holds_for c"
using assms unfolding assert_imp_def by (case_tac c, auto)
lemma test:
assumes "is_final (steps (1, (l, r)) p n1)"
and "is_final (steps (1, (l, r)) p n2)"
shows "Q holds_for (steps (1, (l, r)) p n1) \<longleftrightarrow> Q holds_for (steps (1, (l, r)) p n2)"
proof -
obtain n3 where "n1 = n2 + n3 \<or> n2 = n1 + n3"
by (metis le_iff_add nat_le_linear)
with assms show "Q holds_for (steps (1, (l, r)) p n1) \<longleftrightarrow> Q holds_for (steps (1, (l, r)) p n2)"
by auto
qed
definition
Hoare :: "assert \<Rightarrow> tprog \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50)
where
"{P} p {Q} \<equiv>
(\<forall>l r. P (l, r) \<longrightarrow> (\<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)))"
lemma HoareI:
assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)"
shows "{P} p {Q}"
unfolding Hoare_def using assms by auto
text {*
{P1} A {Q1} {P2} B {Q2} Q1 \<mapsto> P2
-----------------------------------
{P1} A |+| B {Q2}
*}
lemma before_final:
assumes "steps (1, tp) A n = (0, tp')"
obtains n' where "\<not> is_final (steps (1, tp) A n')" and "steps (1, tp) A (Suc n') = (0, tp')"
using assms
apply(induct n)
apply(auto)
by (metis is_final.simps step_red steps.simps steps_0 surj_pair)
lemma t_merge_fetch_pre:
assumes "fetch A s b = (ac, ns)" "s \<le> length A div 2" "tm_wf A" "s \<noteq> 0"
shows "fetch (adjustA |+| B) s b = (ac, if ns = 0 then Suc (length A div 2) else ns)"
using assms
unfolding tm_comp_def
apply(induct A)
apply(auto)
apply(subgoal_tac "2 * (s - Suc 0) < length A \<and> Suc (2 * (s - Suc 0)) < length A")
apply(auto simp: tm_wf_def iseven_def tm_comp_def split: if_splits cell.splits)
oops
lemma t_merge_pre_eq_step:
"\<lbrakk>step (a, b, c) A = cf; tm_wf A; \<not> is_final cf\<rbrakk>
\<Longrightarrow> step (a, b, c) (A |+| B) = cf"
apply(subgoal_tac "a \<le> length A div 2 \<and> a \<noteq> 0")
apply(simp)
apply(case_tac "fetch A a (read c)", simp)
apply(auto)
oops
lemma t_merge_pre_eq:
"\<lbrakk>steps (Suc 0, tp) A stp = cf; \<not> is_final cf; tm_wf A\<rbrakk>
\<Longrightarrow> steps (Suc 0, tp) (A |+| B) stp = cf"
apply(induct stp arbitrary: cf)
apply(auto)[1]
apply(auto)
oops
lemma t_merge_pre_halt_same:
assumes a_ht: "steps (1, tp) A n = (0, tp')"
and a_wf: "t_correct A"
obtains n' where "steps (1, tp) (A |+| B) n' = (Suc (length A div 2), tp')"
proof -
assume a: "\<And>n. steps (1, tp) (A |+| B) n = (Suc (length A div 2), tp') \<Longrightarrow> thesis"
obtain stp' where "\<not> is_final (steps (1, tp) A stp')" and "steps (1, tp) A (Suc stp') = (0, tp')"
using a_ht before_final by blast
then have "steps (1, tp) (A |+| B) (Suc stp') = (Suc (length A div 2), tp')"
sorry (*using a_wf t_merge_pre_halt_same' by blast*)
with a show thesis by blast
qed
lemma steps_comp:
assumes a1: "steps (1, l, r) A n1 = (s1, l1, r1)"
and a2: "steps (1, l1, r1) B n2 = (s2, l2, r2)"
shows "steps (1, l, r) (A |+| B) (n1 + n2) = (s2, l2, r2)"
using assms
apply(induct n2)
apply(simp)
apply(simp add: tm_comp_def)
oops
lemma Hoare_plus:
assumes aimpb: "Q1 \<mapsto> P2"
and A_wf : "tm_wf A"
and B_wf : "tm_wf B"
and A_halt : "{P1} A {Q1}"
and B_halt : "{P2} B {Q2}"
shows "{P1} A |+| B {Q2}"
proof(rule HoareI)
fix l r
assume h: "P1 (l, r)"
then obtain n1 where a: "is_final (steps (1, l, r) A n1)" and b: "Q1 holds_for (steps (1, l, r) A n1)"
using A_halt unfolding Hoare_def by auto
from b aimpb have "P2 holds_for (steps (1, l, r) A n1)"
by (rule holds_for_imp)
then obtain l' r' where "P2 (l', r')"
apply(auto)
apply(case_tac "steps (Suc 0, l, r) A n1")
apply(simp)
done
then obtain n2 where a: "is_final (steps (1, l', r') B n2)" and b: "Q2 holds_for (steps (1, l', r') B n2)"
using B_halt unfolding Hoare_def by auto
show "\<exists>n. is_final (steps (1, l, r) (A |+| B) n) \<and> Q2 holds_for (steps (1, l, r) (A |+| B) n)"
apply(rule_tac x="n1 + n2" in exI)
apply(rule conjI)
apply(simp)
apply(simp only: steps_add[symmetric])
sorry
qed
locale turing_merge =
fixes A :: "tprog" and B :: "tprog" and P1 :: "assert"
and P2 :: "assert"
and P3 :: "assert"
and P4 :: "assert"
and Q1:: "assert"
and Q2 :: "assert"
assumes
A_wf : "tm_wf A"
and B_wf : "tm_wf B"
and A_halt : "P1 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'"
and B_halt : "P2 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) B stp in s = 0 \<and> Q2 tp'"
and A_uhalt : "P3 tp \<Longrightarrow> \<not> (\<exists> stp. is_final (steps (Suc 0, tp) A stp))"
and B_uhalt: "P4 tp \<Longrightarrow> \<not> (\<exists> stp. is_final (steps (Suc 0, tp) B stp))"
begin
text {*
The following lemma tries to derive the Hoare logic rule for sequentially combined TMs.
It deals with the situtation when both @{text "A"} and @{text "B"} are terminated.
*}
lemma t_merge_uhalt_tmp:
assumes B_uh: "\<forall>stp. \<not> is_final (steps (Suc 0, b, c) B stp)"
and merge_ah: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)"
shows "\<forall> stp. \<not> is_final (steps (Suc 0, tp) (A |+| B) stp)"
using B_uh merge_ah
apply(rule_tac allI)
apply(case_tac "stp > stpa")
apply(erule_tac x = "stp - stpa" in allE)
apply(case_tac "(steps (Suc 0, b, c) B (stp - stpa))", simp)
proof -
fix stp a ba ca
assume h1: "\<not> is_final (a, ba, ca)" "stpa < stp"
and h2: "steps (Suc 0, b, c) B (stp - stpa) = (a, ba, ca)"
have "steps (Suc 0 + length A div 2, b, c) (A |+| B) (stp - stpa) =
(if a = 0 then 0 else a + length A div 2, ba, ca)"
using A_wf B_wf h2
by(rule_tac t_merge_snd_eq_steps, auto)
moreover have "a > 0" using h1 by(simp add: is_final_def)
moreover have "\<exists> stpb. stp = stpa + stpb"
using h1 by(rule_tac x = "stp - stpa" in exI, simp)
ultimately show "\<not> is_final (steps (Suc 0, tp) (A |+| B) stp)"
using merge_ah
by(auto simp: steps_add is_final_def)
next
fix stp
assume h: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" "\<not> stpa < stp"
hence "\<exists> stpb. stpa = stp + stpb" apply(rule_tac x = "stpa - stp" in exI, auto) done
thus "\<not> is_final (steps (Suc 0, tp) (A |+| B) stp)"
using h
apply(auto)
apply(cases "steps (Suc 0, tp) (A |+| B) stp", simp add: steps_add is_final_def steps_0)
done
qed
text {*
The following lemma deals with the situation when TM @{text "B"} can not terminate.
*}
lemma t_merge_uhalt:
assumes aimpb: "Q1 \<mapsto> P4"
shows "P1 \<mapsto> \<lambda> tp. \<not> (\<exists> stp. is_final (steps (Suc 0, tp) (A |+| B) stp))"
proof(simp only: assert_imp_def, rule_tac allI, rule_tac impI)
fix tp
assume init_asst: "P1 tp"
show "\<not> (\<exists>stp. is_final (steps (Suc 0, tp) (A |+| B) stp))"
proof -
have "\<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'"
using A_halt[of tp] init_asst
by(simp)
from this obtain stpx where "let (s, tp') = steps (Suc 0, tp) A stpx in s = 0 \<and> Q1 tp'" ..
thus "?thesis"
proof(cases "steps (Suc 0, tp) A stpx", simp, erule_tac conjE)
fix a b c
assume "Q1 (b, c)"
and h3: "steps (Suc 0, tp) A stpx = (0, b, c)"
hence h2: "P4 (b, c)" using aimpb
by(simp add: assert_imp_def)
have "\<exists> stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), b, c)"
using h3 A_wf B_wf
apply(rule_tac stp = stpx in t_merge_pre_halt_same, auto)
done
from this obtain stpa where h4:"steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" ..
have " \<not> (\<exists> stp. is_final (steps (Suc 0, b, c) B stp))"
using B_uhalt [of "(b, c)"] h2 apply simp
done
from this and h4 show "\<forall>stp. \<not> is_final (steps (Suc 0, tp) (A |+| B) stp)"
by(rule_tac t_merge_uhalt_tmp, auto)
qed
qed
qed
end
end