diff -r 8f8db701f69f -r a95987e9c7e9 thys/turing_basic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/turing_basic.thy Sun Jan 13 09:57:28 2013 +0000 @@ -0,0 +1,350 @@ +(* Title: Turing machines + Author: Xu Jian + Maintainer: Xu Jian +*) + +theory turing_basic +imports Main +begin + +section {* Basic definitions of Turing machine *} + +definition + "iseven n \ \x. n = 2 * x" + +datatype action = W0 | W1 | L | R | Nop + +datatype cell = Bk | Oc + +type_synonym tape = "cell list \ cell list" + +type_synonym state = nat + +type_synonym instr = "action \ state" + +type_synonym tprog = "instr list" + +type_synonym config = "state \ 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 \ state \ cell \ instr" +where + "fetch p 0 b = (Nop, 0)" +| "fetch p (Suc s) Bk = + (case nth_of p (2 * s) of + Some i \ i + | None \ (Nop, 0))" +|"fetch p (Suc s) Oc = + (case nth_of p ((2 * s) + 1) of + Some i \ i + | None \ (Nop, 0))" + +fun + update :: "action \ tape \ 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 \ tprog \ config" +where + "step (s, l, r) p = (let (a, s') = fetch p s (read r) in (s', update a (l, r)))" + +fun steps :: "config \ tprog \ nat \ 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 \ bool" +where + "tm_wf p = (length p \ 2 \ iseven (length p) \ (\(a, s) \ set p. s \ length p div 2))" + +(* FIXME: needed? *) +lemma halt_lemma: + "\wf LE; \n. (\ P (f n) \ (f (Suc n), (f n)) \ LE)\ \ \n. P (f n)" +by (metis wf_iff_no_infinite_down_chain) + +abbreviation exponent :: "'a \ nat \ 'a list" ("_ \ _" [100, 99] 100) + where "x \ n == replicate n x" + +definition tinres :: "cell list \ cell list \ bool" + where + "tinres xs ys = (\n. xs = ys @ Bk \ n \ ys = xs @ Bk \ n)" + +fun + shift :: "tprog \ nat \ tprog" +where + "shift p n = (map (\ (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 \ tprog" +where + "adjust p = (map (\ (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 \ tprog \ tprog" ("_ |+| _" [0, 0] 100) +where + "tm_comp p1 p2 = ((adjust p1) @ (shift p2 ((length p1) div 2)))" + +fun + is_final :: "config \ bool" +where + "is_final (s, l, r) = (s = 0)" + +fun + holds_for :: "(tape \ bool) \ config \ 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) + +type_synonym assert = "tape \ bool" + +definition assert_imp :: "assert \ assert \ bool" ("_ \ _" [0, 0] 100) + where + "assert_imp P Q = (\l r. P (l, r) \ Q (l, r))" + +definition + Hoare :: "assert \ tprog \ assert \ bool" ("({(1_)}/ (_)/ {(1_)})" 50) +where + "{P} p {Q} \ + (\l r. P (l, r) \ (\n. is_final (steps (1, (l, r)) p n) \ Q holds_for (steps (1, (l, r)) p n)))" + +lemma HoareI: + assumes "\l r. P (l, r) \ \n. is_final (steps (1, (l, r)) p n) \ Q holds_for (steps (1, (l, r)) p n)" + shows "{P} p {Q}" +unfolding Hoare_def using assms by auto + +lemma HoareI2: + assumes "\l r n. P (l, r) \ is_final (steps (1, (l, r)) p n) \ 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 \ P2 +----------------------------------- + {P1} A |+| B {Q2} +*} + + +lemma Hoare_plus: + assumes aimpb: "Q1 \ 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)" + hence "\n. let (s, tp') = steps (Suc 0, a, b) A n in s = 0 \ 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 \ Q1 tp'" .. + then show "\n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \ s = 0 \ 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 "\ stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \ 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 \ 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 "\ 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 "\ stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" + using g3 A_wf B_wf + sorry + ultimately show "\n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \ s = 0 \ 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_plus: + 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 \ P2" + shows "{P1} A |+| B {Q2}" +unfolding hoare_def +proof(auto split: ) + fix a b + assume h: "P1 (a, b)" + hence "\n. let (s, tp') = steps (Suc 0, a, b) A n in s = 0 \ 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 \ Q1 tp'" .. + then show "\n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \ s = 0 \ 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 "\ stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \ 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 \ 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 "\ 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 "\ stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" + using g3 A_wf B_wf + sorry + ultimately show "\n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \ s = 0 \ 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 \ \ stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \ Q1 tp'" + and B_halt : "P2 tp \ \ stp. let (s, tp') = steps (Suc 0, tp) B stp in s = 0 \ Q2 tp'" + and A_uhalt : "P3 tp \ \ (\ stp. is_final (steps (Suc 0, tp) A stp))" + and B_uhalt: "P4 tp \ \ (\ 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: "\stp. \ 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 "\ stp. \ 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: "\ 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 "\ stpb. stp = stpa + stpb" + using h1 by(rule_tac x = "stp - stpa" in exI, simp) + ultimately show "\ 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)" "\ stpa < stp" + hence "\ stpb. stpa = stp + stpb" apply(rule_tac x = "stpa - stp" in exI, auto) done + thus "\ 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 \ P4" + shows "P1 \ \ tp. \ (\ 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 "\ (\stp. is_final (steps (Suc 0, tp) (A |+| B) stp))" + proof - + have "\ stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \ 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 \ 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 "\ 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 " \ (\ 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 "\stp. \ is_final (steps (Suc 0, tp) (A |+| B) stp)" + by(rule_tac t_merge_uhalt_tmp, auto) + qed + qed +qed +end + +end +