(* 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" + −
+ −
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))"+ −
+ −
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 [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 [simp]: + −
shows "length (adjust p) = length p"+ −
by (simp)+ −
+ −
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 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 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 a b+ −
assume h: "P1 (a, b)"+ −
then have "\<exists>n. is_final (steps (1, a, b) A n) \<and> Q1 holds_for (steps (1, a, b) A n)"+ −
using A_halt unfolding Hoare_def by simp+ −
then obtain n1 where "is_final (steps (1, a, b) A n1) \<and> Q1 holds_for (steps (1, a, b) A stp1)" ..+ −
then show "\<exists>n. is_final (steps (1, a, b) (A |+| B) n) \<and> Q2 holds_for (steps (1, a, b) (A |+| B) n)"+ −
proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE)+ −
fix aa ba c+ −
assume g1: "Q1 (ba, c)" + −
and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)"+ −
hence "P2 (ba, c)"+ −
using aimpb apply(simp add: assert_imp_def)+ −
done+ −
hence "\<exists> stp. is_final (steps (Suc 0, ba, c) B stp) \<and> Q2 holds_for (steps (Suc 0, ba, c) B stp)"+ −
using B_halt unfolding Hoare_def by simp+ −
from this obtain stp2 where + −
"is_final (steps (Suc 0, ba, c) B stp2) \<and> Q2 holds_for (steps (Suc 0, ba, c) B stp2)" ..+ −
thus "?thesis"+ −
proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE)+ −
fix aa bb ca+ −
assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)"+ −
have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)"+ −
using g2 A_wf B_wf+ −
sorry+ −
moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)"+ −
using g3 A_wf B_wf+ −
sorry+ −
ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'"+ −
apply(erule_tac exE, erule_tac exE)+ −
apply(rule_tac x = "stp + stpa" in exI, simp)+ −
using g3 by simp+ −
qed+ −
qed+ −
qed+ −
+ −
lemma Hoare_plus2: + −
assumes A_wf : "tm_wf A"+ −
and B_wf : "tm_wf B"+ −
and A_halt : "{P1} A {Q1}"+ −
and B_halt : "{P2} B {Q2}"+ −
and aimpb: "Q1 \<mapsto> P2"+ −
shows "{P1} A |+| B {Q2}"+ −
unfolding hoare_def+ −
proof(auto split: )+ −
fix a b+ −
assume h: "P1 (a, b)"+ −
hence "\<exists>n. let (s, tp') = steps (Suc 0, a, b) A n in s = 0 \<and> Q1 tp'"+ −
using A_halt unfolding hoare_def by simp+ −
from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \<and> Q1 tp'" ..+ −
then show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'"+ −
proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE)+ −
fix aa ba c+ −
assume g1: "Q1 (ba, c)" + −
and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)"+ −
hence "P2 (ba, c)"+ −
using aimpb apply(simp add: assert_imp_def)+ −
done+ −
hence "\<exists> stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \<and> Q2 tp'"+ −
using B_halt unfolding hoare_def by simp+ −
from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \<and> Q2 tp'" ..+ −
thus "?thesis"+ −
proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE)+ −
fix aa bb ca+ −
assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)"+ −
have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)"+ −
using g2 A_wf B_wf+ −
sorry+ −
moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)"+ −
using g3 A_wf B_wf+ −
sorry+ −
ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'"+ −
apply(erule_tac exE, erule_tac exE)+ −
apply(rule_tac x = "stp + stpa" in exI, simp)+ −
using g3 by simp+ −
qed+ −
qed+ −
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+ −
+ −