diff -r cbb4ac6c8081 -r 1ce04eb1c8ad utm/turing_basic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/utm/turing_basic.thy Sat Sep 29 12:38:12 2012 +0000 @@ -0,0 +1,747 @@ +theory turing_basic +imports Main +begin + +section {* Basic definitions of Turing machine *} + +(* Title: Turing machine's definition and its charater + Author: Xu Jian + Maintainer: Xu Jian +*) + +text {* +\label{description of turing machine} +*} + +section {* Basic definitions of Turing machine *} + +(* Title: Turing machine's definition and its charater + Author: Xu Jian + Maintainer: Xu Jian +*) + +text {* + Actions of Turing machine (Abbreviated TM in the following* ). +*} + +datatype taction = + -- {* Write zero *} + W0 | + -- {* Write one *} + W1 | + -- {* Move left *} + L | + -- {* Move right *} + R | + -- {* Do nothing *} + Nop + +text {* + Tape contents in every block. +*} + +datatype block = + -- {* Blank *} + Bk | + -- {* Occupied *} + Oc + +text {* + Tape is represented as a pair of lists $(L_{left}, L_{right})$, + where $L_left$, named {\em left list}, is used to represent + the tape to the left of RW-head and + $L_{right}$, named {\em right list}, is used to represent the tape + under and to the right of RW-head. +*} + +type_synonym tape = "block list \ block list" + +text {* The state of turing machine.*} +type_synonym tstate = nat + +text {* + Turing machine instruction is represented as a + pair @{text "(action, next_state)"}, + where @{text "action"} is the action to take at the current state + and @{text "next_state"} is the next state the machine is getting into + after the action. +*} +type_synonym tinst = "taction \ tstate" + +text {* + Program of Turing machine is represented as a list of Turing instructions + and the execution of the program starts from the head of the list. + *} +type_synonym tprog = "tinst list" + + +text {* + Turing machine configuration, which consists of the current state + and the tape. +*} +type_synonym t_conf = "tstate \ tape" + +fun nth_of :: "'a list \ nat \ 'a option" + where + "nth_of xs n = (if n < length xs then Some (xs!n) + else None)" + +text {* + The function used to fetech instruction out of Turing program. + *} + +fun fetch :: "tprog \ tstate \ block \ tinst" + where + "fetch p s b = (if s = 0 then (Nop, 0) else + case b of + Bk \ case nth_of p (2 * (s - 1)) of + Some i \ i + | None \ (Nop, 0) + | Oc \ case nth_of p (2 * (s - 1) +1) of + Some i \ i + | None \ (Nop, 0))" + + +fun new_tape :: "taction \ tape \ tape" +where + "new_tape action (leftn, rightn) = (case action of + W0 \ (leftn, Bk#(tl rightn)) | + W1 \ (leftn, Oc#(tl rightn)) | + L \ (if leftn = [] then (tl leftn, Bk#rightn) + else (tl leftn, (hd leftn) # rightn)) | + R \ if rightn = [] then (Bk#leftn,tl rightn) + else ((hd rightn)#leftn, tl rightn) | + Nop \ (leftn, rightn) + )" + +text {* + The one step function used to transfer Turing machine configuration. +*} +fun tstep :: "t_conf \ tprog \ t_conf" + where + "tstep c p = (let (s, l, r) = c in + let (ac, ns) = (fetch p s (case r of [] \ Bk | + x # xs \ x)) in + (ns, new_tape ac (l, r)))" + +text {* + The many-step function. +*} +fun steps :: "t_conf \ tprog \ nat \ t_conf" + where + "steps c p 0 = c" | + "steps c p (Suc n) = steps (tstep c p) p n" + +lemma tstep_red: "steps c p (Suc n) = tstep (steps c p n) p" +proof(induct n arbitrary: c) + fix c + show "steps c p (Suc 0) = tstep (steps c p 0) p" by(simp add: steps.simps) +next + fix n c + assume ind: "\ c. steps c p (Suc n) = tstep (steps c p n) p" + have "steps (tstep c p) p (Suc n) = tstep (steps (tstep c p) p n) p" + by(rule ind) + thus "steps c p (Suc (Suc n)) = tstep (steps c p (Suc n)) p" by(simp add: steps.simps) +qed + +declare Let_def[simp] option.split[split] + +definition + "iseven n \ \ x. n = 2 * x" + + +text {* + The following @{text "t_correct"} function is used to specify the wellformedness of Turing + machine. +*} +fun t_correct :: "tprog \ bool" + where + "t_correct p = (length p \ 2 \ iseven (length p) \ + list_all (\ (acn, s). s \ length p div 2) p)" + +declare t_correct.simps[simp del] + +lemma allimp: "\\x. P x \ Q x; \x. P x\ \ \x. Q x" +by(auto elim: allE) + +lemma halt_lemma: "\wf LE; \ n. (\ P (f n) \ (f (Suc n), (f n)) \ LE)\ \ \ n. P (f n)" +apply(rule exCI, drule allimp, auto) +apply(drule_tac f = f in wf_inv_image, simp add: inv_image_def) +apply(erule wf_induct, auto) +done + +lemma steps_add: "steps c t (x + y) = steps (steps c t x) t y" +by(induct x arbitrary: c, auto simp: steps.simps tstep_red) + +lemma listall_set: "list_all p t \ \ a \ set t. p a" +by(induct t, auto) + +lemma fetch_ex: "\b a. fetch T aa ab = (b, a)" +by(simp add: fetch.simps) +definition exponent :: "'a \ nat \ 'a list" ("_\<^bsup>_\<^esup>" [0, 0]100) + where "exponent x n = replicate n x" + +text {* + @{text "tinres l1 l2"} means left list @{text "l1"} is congruent with + @{text "l2"} with respect to the execution of Turing machine. + Appending Blank to the right of eigther one does not affect the + outcome of excution. +*} + +definition tinres :: "block list \ block list \ bool" + where + "tinres bx by = (\ n. bx = by@Bk\<^bsup>n\<^esup> \ by = bx @ Bk\<^bsup>n\<^esup>)" + +lemma exp_zero: "a\<^bsup>0\<^esup> = []" +by(simp add: exponent_def) +lemma exp_ind_def: "a\<^bsup>Suc x \<^esup> = a # a\<^bsup>x\<^esup>" +by(simp add: exponent_def) + +text {* + The following lemma shows the meaning of @{text "tinres"} with respect to + one step execution. + *} +lemma tinres_step: + "\tinres l l'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l', r) t = (sb, lb, rb)\ + \ tinres la lb \ ra = rb \ sa = sb" +apply(auto simp: tstep.simps fetch.simps new_tape.simps + split: if_splits taction.splits list.splits + block.splits) +apply(case_tac [!] "t ! (2 * (ss - Suc 0))", + auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits + block.splits) +apply(case_tac [!] "t ! (2 * (ss - Suc 0) + Suc 0)", + auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits + block.splits) +done + +declare tstep.simps[simp del] steps.simps[simp del] + +text {* + The following lemma shows the meaning of @{text "tinres"} with respect to + many step execution. + *} +lemma tinres_steps: + "\tinres l l'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l', r) t stp = (sb, lb, rb)\ + \ tinres la lb \ ra = rb \ sa = sb" +apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps) +apply(simp add: tstep_red) +apply(case_tac "(steps (ss, l, r) t stp)") +apply(case_tac "(steps (ss, l', r) t stp)") +proof - + fix stp sa la ra sb lb rb a b c aa ba ca + assume ind: "\sa la ra sb lb rb. \steps (ss, l, r) t stp = (sa, la, ra); + steps (ss, l', r) t stp = (sb, lb, rb)\ \ tinres la lb \ ra = rb \ sa = sb" + and h: " tinres l l'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)" + "tstep (steps (ss, l', r) t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" + "steps (ss, l', r) t stp = (aa, ba, ca)" + have "tinres b ba \ c = ca \ a = aa" + apply(rule_tac ind, simp_all add: h) + done + thus "tinres la lb \ ra = rb \ sa = sb" + apply(rule_tac l = b and l' = ba and r = c and ss = a + and t = t in tinres_step) + using h + apply(simp, simp, simp) + done +qed + +text {* + The following function @{text "tshift tp n"} is used to shift Turing programs + @{text "tp"} by @{text "n"} when it is going to be combined with others. + *} + +fun tshift :: "tprog \ nat \ tprog" + where + "tshift tp off = (map (\ (action, state). (action, (if state = 0 then 0 + else state + off))) tp)" + +text {* + When two Turing programs are combined, the end state (state @{text "0"}) of the one + at the prefix position needs to be connected to the start state + of the one at postfix position. If @{text "tp"} is the Turing program + to be at the prefix, @{text "change_termi_state tp"} is the transformed Turing program. + *} +fun change_termi_state :: "tprog \ tprog" + where + "change_termi_state t = + (map (\ (acn, ns). if ns = 0 then (acn, Suc ((length t) div 2)) else (acn, ns)) t)" + +text {* + @{text "t_add tp1 tp2"} is the combined Truing program. +*} + +fun t_add :: "tprog \ tprog \ tprog" ("_ |+| _" [0, 0] 100) + where + "t_add t1 t2 = ((change_termi_state t1) @ (tshift t2 ((length t1) div 2)))" + +text {* + Tests whether the current configuration is at state @{text "0"}. +*} +definition isS0 :: "t_conf \ bool" + where + "isS0 c = (let (s, l, r) = c in s = 0)" + +declare tstep.simps[simp del] steps.simps[simp del] + t_add.simps[simp del] fetch.simps[simp del] + new_tape.simps[simp del] + + +text {* + Single step execution starting from state @{text "0"} will not make any progress. +*} +lemma tstep_0: "tstep (0, tp) p = (0, tp)" +apply(simp add: tstep.simps fetch.simps new_tape.simps) +done + + +text {* + Many step executions starting from state @{text "0"} will not make any progress. +*} + +lemma steps_0: "steps (0, tp) p stp = (0, tp)" +apply(induct stp) +apply(simp add: steps.simps) +apply(simp add: tstep_red tstep_0) +done + +lemma s_keep_step: "\a \ length A div 2; tstep (a, b, c) A = (s, l, r); t_correct A\ + \ s \ length A div 2" +apply(simp add: tstep.simps fetch.simps t_correct.simps iseven_def + split: if_splits block.splits list.splits) +apply(case_tac [!] a, auto simp: list_all_length) +apply(erule_tac x = "2 * nat" in allE, auto) +apply(erule_tac x = "2 * nat" in allE, auto) +apply(erule_tac x = "Suc (2 * nat)" in allE, auto) +done + +lemma s_keep: "\steps (Suc 0, tp) A stp = (s, l, r); t_correct A\ \ s \ length A div 2" +proof(induct stp arbitrary: s l r) + case 0 thus "?case" by(auto simp: t_correct.simps steps.simps) +next + fix stp s l r + assume ind: "\s l r. \steps (Suc 0, tp) A stp = (s, l, r); t_correct A\ \ s \ length A div 2" + and h1: "steps (Suc 0, tp) A (Suc stp) = (s, l, r)" + and h2: "t_correct A" + from h1 h2 show "s \ length A div 2" + proof(simp add: tstep_red, cases "(steps (Suc 0, tp) A stp)", simp) + fix a b c + assume h3: "tstep (a, b, c) A = (s, l, r)" + and h4: "steps (Suc 0, tp) A stp = (a, b, c)" + have "a \ length A div 2" + using h2 h4 + by(rule_tac l = b and r = c in ind, auto) + thus "?thesis" + using h3 h2 + by(simp add: s_keep_step) + qed +qed + +lemma t_merge_fetch_pre: + "\fetch A s b = (ac, ns); s \ length A div 2; t_correct A; s \ 0\ \ + fetch (A |+| B) s b = (ac, if ns = 0 then Suc (length A div 2) else ns)" +apply(subgoal_tac "2 * (s - Suc 0) < length A \ Suc (2 * (s - Suc 0)) < length A") +apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits) +apply(simp_all add: nth_append change_termi_state.simps) +done + +lemma [simp]: "\\ a \ length A div 2; t_correct A\ \ fetch A a b = (Nop, 0)" +apply(auto simp: fetch.simps del: nth_of.simps split: block.splits) +apply(case_tac [!] a, auto simp: t_correct.simps iseven_def) +done + +lemma [elim]: "\t_correct A; \ isS0 (tstep (a, b, c) A)\ \ a \ length A div 2" +apply(rule_tac classical, auto simp: tstep.simps new_tape.simps isS0_def) +done + +lemma [elim]: "\t_correct A; \ isS0 (tstep (a, b, c) A)\ \ 0 < a" +apply(rule_tac classical, simp add: tstep_0 isS0_def) +done + + +lemma t_merge_pre_eq_step: "\tstep (a, b, c) A = cf; t_correct A; \ isS0 cf\ + \ tstep (a, b, c) (A |+| B) = cf" +apply(subgoal_tac "a \ length A div 2 \ a \ 0") +apply(simp add: tstep.simps) +apply(case_tac "fetch A a (case c of [] \ Bk | x # xs \ x)", simp) +apply(drule_tac B = B in t_merge_fetch_pre, simp, simp, simp, simp add: isS0_def, auto) +done + +lemma t_merge_pre_eq: "\steps (Suc 0, tp) A stp = cf; \ isS0 cf; t_correct A\ + \ steps (Suc 0, tp) (A |+| B) stp = cf" +proof(induct stp arbitrary: cf) + case 0 thus "?case" by(simp add: steps.simps) +next + fix stp cf + assume ind: "\cf. \steps (Suc 0, tp) A stp = cf; \ isS0 cf; t_correct A\ + \ steps (Suc 0, tp) (A |+| B) stp = cf" + and h1: "steps (Suc 0, tp) A (Suc stp) = cf" + and h2: "\ isS0 cf" + and h3: "t_correct A" + from h1 h2 h3 show "steps (Suc 0, tp) (A |+| B) (Suc stp) = cf" + proof(simp add: tstep_red, cases "steps (Suc 0, tp) (A) stp", simp) + fix a b c + assume h4: "tstep (a, b, c) A = cf" + and h5: "steps (Suc 0, tp) A stp = (a, b, c)" + have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)" + proof(cases a) + case 0 thus "?thesis" + using h4 h2 + apply(simp add: tstep_0, cases cf, simp add: isS0_def) + done + next + case (Suc n) thus "?thesis" + using h5 h3 + apply(rule_tac ind, auto simp: isS0_def) + done + qed + thus "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = cf" + using h4 h5 h3 h2 + apply(simp) + apply(rule t_merge_pre_eq_step, auto) + done + qed +qed + +declare nth.simps[simp del] tshift.simps[simp del] change_termi_state.simps[simp del] + +lemma [simp]: "length (change_termi_state A) = length A" +by(simp add: change_termi_state.simps) + +lemma first_halt_point: "steps (Suc 0, tp) A stp = (0, tp') + \ \stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" +proof(induct stp) + case 0 thus "?case" by(simp add: steps.simps) +next + case (Suc n) + fix stp + assume ind: "steps (Suc 0, tp) A stp = (0, tp') \ + \stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + and h: "steps (Suc 0, tp) A (Suc stp) = (0, tp')" + from h show "\stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp, case_tac a) + fix a b c + assume g1: "a = (0::nat)" + and g2: "tstep (a, b, c) A = (0, tp')" + and g3: "steps (Suc 0, tp) A stp = (a, b, c)" + have "steps (Suc 0, tp) A stp = (0, tp')" + using g2 g1 g3 + by(simp add: tstep_0) + hence "\ stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + by(rule ind) + thus "\stp. \ isS0 (steps (Suc 0, tp) A stp) \ tstep (steps (Suc 0, tp) A stp) A = (0, tp')" + apply(simp add: tstep_red) + done + next + fix a b c nat + assume g1: "steps (Suc 0, tp) A stp = (a, b, c)" + and g2: "steps (Suc 0, tp) A (Suc stp) = (0, tp')" "a= Suc nat" + thus "\stp. \ isS0 (steps (Suc 0, tp) A stp) \ tstep (steps (Suc 0, tp) A stp) A = (0, tp')" + apply(rule_tac x = stp in exI) + apply(simp add: isS0_def tstep_red) + done + qed +qed + +lemma t_merge_pre_halt_same': + "\\ isS0 (steps (Suc 0, tp) A stp) ; steps (Suc 0, tp) A (Suc stp) = (0, tp'); t_correct A\ + \ steps (Suc 0, tp) (A |+| B) (Suc stp) = (Suc (length A div 2), tp')" +proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp) + fix a b c + assume h1: "\ isS0 (a, b, c)" + and h2: "tstep (a, b, c) A = (0, tp')" + and h3: "t_correct A" + and h4: "steps (Suc 0, tp) A stp = (a, b, c)" + have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)" + using h1 h4 h3 + apply(rule_tac t_merge_pre_eq, auto) + done + moreover have "tstep (a, b, c) (A |+| B) = (Suc (length A div 2), tp')" + using h2 h3 h1 h4 + apply(simp add: tstep.simps) + apply(case_tac " fetch A a (case c of [] \ Bk | x # xs \ x)", simp) + apply(drule_tac B = B in t_merge_fetch_pre, auto simp: isS0_def intro: s_keep) + done + ultimately show "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = (Suc (length A div 2), tp')" + by(simp) +qed + +text {* + When Turing machine @{text "A"} and @{text "B"} are combined and the execution + of @{text "A"} can termination within @{text "stp"} steps, + the combined machine @{text "A |+| B"} will eventually get into the starting + state of machine @{text "B"}. +*} +lemma t_merge_pre_halt_same: " + \steps (Suc 0, tp) A stp = (0, tp'); t_correct A; t_correct B\ + \ \ stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), tp')" +proof - + assume a_wf: "t_correct A" + and b_wf: "t_correct B" + and a_ht: "steps (Suc 0, tp) A stp = (0, tp')" + have halt_point: "\ stp. \ isS0 (steps (Suc 0, tp) A stp) \ steps (Suc 0, tp) A (Suc stp) = (0, tp')" + using a_ht + by(erule_tac first_halt_point) + then obtain stp' where "\ isS0 (steps (Suc 0, tp) A stp') \ steps (Suc 0, tp) A (Suc stp') = (0, tp')".. + hence "steps (Suc 0, tp) (A |+| B) (Suc stp') = (Suc (length A div 2), tp')" + using a_wf + apply(rule_tac t_merge_pre_halt_same', auto) + done + thus "?thesis" .. +qed + +lemma fetch_0: "fetch p 0 b = (Nop, 0)" +by(simp add: fetch.simps) + +lemma [simp]: "length (tshift B x) = length B" +by(simp add: tshift.simps) + +lemma [simp]: "t_correct A \ 2 * (length A div 2) = length A" +apply(simp add: t_correct.simps iseven_def, auto) +done + +lemma t_merge_fetch_snd: + "\fetch B a b = (ac, ns); t_correct A; t_correct B; a > 0 \ + \ fetch (A |+| B) (a + length A div 2) b + = (ac, if ns = 0 then 0 else ns + length A div 2)" +apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits) +apply(case_tac [!] a, simp_all) +apply(simp_all add: nth_append change_termi_state.simps tshift.simps) +done + +lemma t_merge_snd_eq_step: + "\tstep (s, l, r) B = (s', l', r'); t_correct A; t_correct B; s > 0\ + \ tstep (s + length A div 2, l, r) (A |+| B) = + (if s' = 0 then 0 else s' + length A div 2, l' ,r') " +apply(simp add: tstep.simps) +apply(cases "fetch B s (case r of [] \ Bk | x # xs \ x)") +apply(auto simp: t_merge_fetch_snd) +apply(frule_tac [!] t_merge_fetch_snd, auto) +done + +text {* + Relates the executions of TM @{text "B"}, one is when @{text "B"} is executed alone, + the other is the execution when @{text "B"} is in the combined TM. +*} +lemma t_merge_snd_eq_steps: + "\t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); s > 0\ + \ steps (s + length A div 2, l, r) (A |+| B) stp = + (if s' = 0 then 0 else s' + length A div 2, l', r')" +proof(induct stp arbitrary: s' l' r') + case 0 thus "?case" + by(simp add: steps.simps) +next + fix stp s' l' r' + assume ind: "\s' l' r'. \t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); 0 < s\ + \ steps (s + length A div 2, l, r) (A |+| B) stp = + (if s' = 0 then 0 else s' + length A div 2, l', r')" + and h1: "steps (s, l, r) B (Suc stp) = (s', l', r')" + and h2: "t_correct A" + and h3: "t_correct B" + and h4: "0 < s" + from h1 show "steps (s + length A div 2, l, r) (A |+| B) (Suc stp) + = (if s' = 0 then 0 else s' + length A div 2, l', r')" + proof(simp only: tstep_red, cases "steps (s, l, r) B stp") + fix a b c + assume h5: "steps (s, l, r) B stp = (a, b, c)" "tstep (steps (s, l, r) B stp) B = (s', l', r')" + hence h6: "(steps (s + length A div 2, l, r) (A |+| B) stp) = + ((if a = 0 then 0 else a + length A div 2, b, c))" + using h2 h3 h4 + by(rule_tac ind, auto) + thus "tstep (steps (s + length A div 2, l, r) (A |+| B) stp) (A |+| B) = + (if s' = 0 then 0 else s'+ length A div 2, l', r')" + using h5 + proof(auto) + assume "tstep (0, b, c) B = (0, l', r')" thus "tstep (0, b, c) (A |+| B) = (0, l', r')" + by(simp add: tstep_0) + next + assume "tstep (0, b, c) B = (s', l', r')" "0 < s'" + thus "tstep (0, b, c) (A |+| B) = (s' + length A div 2, l', r')" + by(simp add: tstep_0) + next + assume "tstep (a, b, c) B = (0, l', r')" "0 < a" + thus "tstep (a + length A div 2, b, c) (A |+| B) = (0, l', r')" + using h2 h3 + by(drule_tac t_merge_snd_eq_step, auto) + next + assume "tstep (a, b, c) B = (s', l', r')" "0 < a" "0 < s'" + thus "tstep (a + length A div 2, b, c) (A |+| B) = (s' + length A div 2, l', r')" + using h2 h3 + by(drule_tac t_merge_snd_eq_step, auto) + qed + qed +qed + +lemma t_merge_snd_halt_eq: + "\steps (Suc 0, tp) B stp = (0, tp'); t_correct A; t_correct B\ + \ \stp. steps (Suc (length A div 2), tp) (A |+| B) stp = (0, tp')" +apply(case_tac tp, cases tp', simp) +apply(drule_tac s = "Suc 0" in t_merge_snd_eq_steps, auto) +done + +lemma t_inj: "\steps (Suc 0, tp) A stpa = (0, tp1); steps (Suc 0, tp) A stpb = (0, tp2)\ + \ tp1 = tp2" +proof - + assume h1: "steps (Suc 0, tp) A stpa = (0, tp1)" + and h2: "steps (Suc 0, tp) A stpb = (0, tp2)" + thus "?thesis" + proof(cases "stpa < stpb") + case True thus "?thesis" + using h1 h2 + apply(drule_tac less_imp_Suc_add, auto) + apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0) + done + next + case False thus "?thesis" + using h1 h2 + apply(drule_tac leI) + apply(case_tac "stpb = stpa", auto) + apply(subgoal_tac "stpb < stpa") + apply(drule_tac less_imp_Suc_add, auto) + apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0) + done + qed +qed + +type_synonym t_assert = "tape \ bool" + +definition t_imply :: "t_assert \ t_assert \ bool" ("_ \-> _" [0, 0] 100) + where + "t_imply a1 a2 = (\ tp. a1 tp \ a2 tp)" + + +locale turing_merge = + fixes A :: "tprog" and B :: "tprog" and P1 :: "t_assert" + and P2 :: "t_assert" + and P3 :: "t_assert" + and P4 :: "t_assert" + and Q1:: "t_assert" + and Q2 :: "t_assert" + assumes + A_wf : "t_correct A" + and B_wf : "t_correct 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. isS0 (steps (Suc 0, tp) A stp))" + and B_uhalt: "P4 tp \ \ (\ stp. isS0 (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_halt: + assumes aimpb: "Q1 \-> P2" + shows "P1 \-> \ tp. (\ stp tp'. steps (Suc 0, tp) (A |+| B) stp = (0, tp') \ Q2 tp')" +proof(simp add: t_imply_def, auto) + fix a b + assume h: "P1 (a, b)" + hence "\ stp. let (s, tp') = steps (Suc 0, a, b) A stp in s = 0 \ Q1 tp'" + using A_halt by simp + from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \ Q1 tp'" .. + thus "\stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \ Q2 (aa, ba)" + 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: t_imply_def) + done + hence "\ stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \ Q2 tp'" + using B_halt 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 + by(rule_tac t_merge_pre_halt_same, auto) + moreover have "\ stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" + using g3 A_wf B_wf + apply(rule_tac t_merge_snd_halt_eq, auto) + done + ultimately show "\stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \ Q2 (aa, ba)" + apply(erule_tac exE, erule_tac exE) + apply(rule_tac x = "stp + stpa" in exI, simp add: steps_add) + using g3 by simp + qed + qed +qed + +lemma t_merge_uhalt_tmp: + assumes B_uh: "\stp. \ isS0 (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. \ isS0 (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: "\ isS0 (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: isS0_def) + moreover have "\ stpb. stp = stpa + stpb" + using h1 by(rule_tac x = "stp - stpa" in exI, simp) + ultimately show "\ isS0 (steps (Suc 0, tp) (A |+| B) stp)" + using merge_ah + by(auto simp: steps_add isS0_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 "\ isS0 (steps (Suc 0, tp) (A |+| B) stp)" + using h + apply(auto) + apply(cases "steps (Suc 0, tp) (A |+| B) stp", simp add: steps_add isS0_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. isS0 (steps (Suc 0, tp) (A |+| B) stp))" +proof(simp only: t_imply_def, rule_tac allI, rule_tac impI) + fix tp + assume init_asst: "P1 tp" + show "\ (\stp. isS0 (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: t_imply_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. isS0 (steps (Suc 0, b, c) B stp))" + using B_uhalt [of "(b, c)"] h2 apply simp + done + from this and h4 show "\stp. \ isS0 (steps (Suc 0, tp) (A |+| B) stp)" + by(rule_tac t_merge_uhalt_tmp, auto) + qed + qed +qed +end + +end +