diff -r 000000000000 -r aa8656a8dbef abacus.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/abacus.thy Mon Dec 24 01:26:23 2012 +0000 @@ -0,0 +1,6923 @@ +header {* + {\em abacus} a kind of register machine +*} + +theory abacus +imports Main turing_basic +begin + +text {* + {\em Abacus} instructions: +*} + +datatype abc_inst = + -- {* @{text "Inc n"} increments the memory cell (or register) with address @{text "n"} by one. + *} + Inc nat + -- {* + @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. + If cell @{text "n"} is already zero, no decrements happens and the executio jumps to + the instruction labeled by @{text "label"}. + *} + | Dec nat nat + -- {* + @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}. + *} + | Goto nat + + +text {* + Abacus programs are defined as lists of Abacus instructions. +*} +type_synonym abc_prog = "abc_inst list" + +section {* + Sample Abacus programs + *} + +text {* + Abacus for addition and clearance. +*} +fun plus_clear :: "nat \ nat \ nat \ abc_prog" + where + "plus_clear m n e = [Dec m e, Inc n, Goto 0]" + +text {* + Abacus for clearing memory untis. +*} +fun clear :: "nat \ nat \ abc_prog" + where + "clear n e = [Dec n e, Goto 0]" + +fun plus:: "nat \ nat \ nat \ nat \ abc_prog" + where + "plus m n p e = [Dec m 4, Inc n, Inc p, + Goto 0, Dec p e, Inc m, Goto 4]" + +fun mult :: "nat \ nat \ nat \ nat \ nat \ abc_prog" + where + "mult m1 m2 n p e = [Dec m1 e]@ plus m1 m2 p 1" + +fun expo :: "nat \ nat \ nat \ nat \ nat \ abc_prog" + where + "expo n m1 m2 p e = [Inc n, Dec m1 e] @ mult m2 n n p 2" + + +text {* + The state of Abacus machine. + *} +type_synonym abc_state = nat + +(* text {* + The memory of Abacus machine is defined as a function from address to contents. +*} +type_synonym abc_mem = "nat \ nat" *) + +text {* + The memory of Abacus machine is defined as a list of contents, with + every units addressed by index into the list. + *} +type_synonym abc_lm = "nat list" + +text {* + Fetching contents out of memory. Units not represented by list elements are considered + as having content @{text "0"}. +*} +fun abc_lm_v :: "abc_lm \ nat \ nat" + where + "abc_lm_v lm n = (if (n < length lm) then (lm!n) else 0)" + +(* +fun abc_l2m :: "abc_lm \ abc_mem" + where + "abc_l2m lm = abc_lm_v lm" +*) + +text {* + Set the content of memory unit @{text "n"} to value @{text "v"}. + @{text "am"} is the Abacus memory before setting. + If address @{text "n"} is outside to scope of @{text "am"}, @{text "am"} + is extended so that @{text "n"} becomes in scope. +*} +fun abc_lm_s :: "abc_lm \ nat \ nat \ abc_lm" + where + "abc_lm_s am n v = (if (n < length am) then (am[n:=v]) else + am@ (replicate (n - length am) 0) @ [v])" + + +text {* + The configuration of Abaucs machines consists of its current state and its + current memory: +*} +type_synonym abc_conf_l = "abc_state \ abc_lm" + +text {* + Fetch instruction out of Abacus program: +*} + +fun abc_fetch :: "nat \ abc_prog \ abc_inst option" + where + "abc_fetch s p = (if (s < length p) then Some (p ! s) + else None)" + +text {* + Single step execution of Abacus machine. If no instruction is feteched, + configuration does not change. +*} +fun abc_step_l :: "abc_conf_l \ abc_inst option \ abc_conf_l" + where + "abc_step_l (s, lm) a = (case a of + None \ (s, lm) | + Some (Inc n) \ (let nv = abc_lm_v lm n in + (s + 1, abc_lm_s lm n (nv + 1))) | + Some (Dec n e) \ (let nv = abc_lm_v lm n in + if (nv = 0) then (e, abc_lm_s lm n 0) + else (s + 1, abc_lm_s lm n (nv - 1))) | + Some (Goto n) \ (n, lm) + )" + +text {* + Multi-step execution of Abacus machine. +*} +fun abc_steps_l :: "abc_conf_l \ abc_prog \ nat \ abc_conf_l" + where + "abc_steps_l (s, lm) p 0 = (s, lm)" | + "abc_steps_l (s, lm) p (Suc n) = abc_steps_l (abc_step_l (s, lm) (abc_fetch s p)) p n" + +section {* + Compiling Abacus machines into Truing machines +*} + + +subsection {* + Compiling functions +*} + +text {* + @{text "findnth n"} returns the TM which locates the represention of + memory cell @{text "n"} on the tape and changes representation of zero + on the way. +*} + +fun findnth :: "nat \ tprog" + where + "findnth 0 = []" | + "findnth (Suc n) = (findnth n @ [(W1, 2 * n + 1), + (R, 2 * n + 2), (R, 2 * n + 3), (R, 2 * n + 2)])" + +text {* + @{text "tinc_b"} returns the TM which increments the representation + of the memory cell under rw-head by one and move the representation + of cells afterwards to the right accordingly. + *} + +definition tinc_b :: "tprog" + where + "tinc_b \ [(W1, 1), (R, 2), (W1, 3), (R, 2), (W1, 3), (R, 4), + (L, 7), (W0, 5), (R, 6), (W0, 5), (W1, 3), (R, 6), + (L, 8), (L, 7), (R, 9), (L, 7), (R, 10), (W0, 9)]" + +text {* + @{text "tshift tm off"} shifts @{text "tm"} by offset @{text "off"}, leaving + instructions concerning state @{text "0"} unchanged, because state @{text "0"} + is the end state, which needs not be changed with shift operation. + *} + +fun tshift :: "tprog \ nat \ tprog" + where + "tshift tp off = (map (\ (action, state). + (action, (if state = 0 then 0 + else state + off))) tp)" + +text {* + @{text "tinc ss n"} returns the TM which simulates the execution of + Abacus instruction @{text "Inc n"}, assuming that TM is located at + location @{text "ss"} in the final TM complied from the whole + Abacus program. +*} + +fun tinc :: "nat \ nat \ tprog" + where + "tinc ss n = tshift (findnth n @ tshift tinc_b (2 * n)) (ss - 1)" + +text {* + @{text "tinc_b"} returns the TM which decrements the representation + of the memory cell under rw-head by one and move the representation + of cells afterwards to the left accordingly. + *} + +definition tdec_b :: "tprog" + where + "tdec_b \ [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3), + (R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8), + (L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9), + (R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11), + (R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14), + (R, 0), (W0, 16)]" + +text {* + @{text "sete tp e"} attaches the termination edges (edges leading to state @{text "0"}) + of TM @{text "tp"} to the intruction labelled by @{text "e"}. + *} + +fun sete :: "tprog \ nat \ tprog" + where + "sete tp e = map (\ (action, state). (action, if state = 0 then e else state)) tp" + +text {* + @{text "tdec ss n label"} returns the TM which simulates the execution of + Abacus instruction @{text "Dec n label"}, assuming that TM is located at + location @{text "ss"} in the final TM complied from the whole + Abacus program. +*} + +fun tdec :: "nat \ nat \ nat \ tprog" + where + "tdec ss n e = sete (tshift (findnth n @ tshift tdec_b (2 * n)) + (ss - 1)) e" + +text {* + @{text "tgoto f(label)"} returns the TM simulating the execution of Abacus instruction + @{text "Goto label"}, where @{text "f(label)"} is the corresponding location of + @{text "label"} in the final TM compiled from the overall Abacus program. +*} + +fun tgoto :: "nat \ tprog" + where + "tgoto n = [(Nop, n), (Nop, n)]" + +text {* + The layout of the final TM compiled from an Abacus program is represented + as a list of natural numbers, where the list element at index @{text "n"} represents the + starting state of the TM simulating the execution of @{text "n"}-th instruction + in the Abacus program. +*} + +type_synonym layout = "nat list" + +text {* + @{text "length_of i"} is the length of the + TM simulating the Abacus instruction @{text "i"}. +*} +fun length_of :: "abc_inst \ nat" + where + "length_of i = (case i of + Inc n \ 2 * n + 9 | + Dec n e \ 2 * n + 16 | + Goto n \ 1)" + +text {* + @{text "layout_of ap"} returns the layout of Abacus program @{text "ap"}. +*} +fun layout_of :: "abc_prog \ layout" + where "layout_of ap = map length_of ap" + + +text {* + @{text "start_of layout n"} looks out the starting state of @{text "n"}-th + TM in the finall TM. +*} + +fun start_of :: "nat list \ nat \ nat" + where + "start_of ly 0 = Suc 0" | + "start_of ly (Suc as) = + (if as < length ly then start_of ly as + (ly ! as) + else start_of ly as)" + +text {* + @{text "ci lo ss i"} complies Abacus instruction @{text "i"} + assuming the TM of @{text "i"} starts from state @{text "ss"} + within the overal layout @{text "lo"}. +*} + +fun ci :: "layout \ nat \ abc_inst \ tprog" + where + "ci ly ss i = (case i of + Inc n \ tinc ss n | + Dec n e \ tdec ss n (start_of ly e) | + Goto n \ tgoto (start_of ly n))" + +text {* + @{text "tpairs_of ap"} transfroms Abacus program @{text "ap"} pairing + every instruction with its starting state. +*} +fun tpairs_of :: "abc_prog \ (nat \ abc_inst) list" + where "tpairs_of ap = (zip (map (start_of (layout_of ap)) + [0..<(length ap)]) ap)" + + +text {* + @{text "tms_of ap"} returns the list of TMs, where every one of them simulates + the corresponding Abacus intruction in @{text "ap"}. +*} + +fun tms_of :: "abc_prog \ tprog list" + where "tms_of ap = map (\ (n, tm). ci (layout_of ap) n tm) + (tpairs_of ap)" + +text {* + @{text "tm_of ap"} returns the final TM machine compiled from Abacus program @{text "ap"}. +*} +fun tm_of :: "abc_prog \ tprog" + where "tm_of ap = concat (tms_of ap)" + +text {* + The following two functions specify the well-formedness of complied TM. +*} +fun t_ncorrect :: "tprog \ bool" + where + "t_ncorrect tp = (length tp mod 2 = 0)" + +fun abc2t_correct :: "abc_prog \ bool" + where + "abc2t_correct ap = list_all (\ (n, tm). + t_ncorrect (ci (layout_of ap) n tm)) (tpairs_of ap)" + +lemma findnth_length: "length (findnth n) div 2 = 2 * n" +apply(induct n, simp, simp) +done + +lemma ci_length : "length (ci ns n ai) div 2 = length_of ai" +apply(auto simp: ci.simps tinc_b_def tdec_b_def findnth_length + split: abc_inst.splits) +done + +subsection {* + Representation of Abacus memory by TM tape +*} + +consts tape_of :: "'a \ block list" ("<_>" 100) + +text {* + @{text "tape_of_nat_list am"} returns the TM tape representing + Abacus memory @{text "am"}. + *} + +fun tape_of_nat_list :: "nat list \ block list" + where + "tape_of_nat_list [] = []" | + "tape_of_nat_list [n] = Oc\<^bsup>n+1\<^esup>" | + "tape_of_nat_list (n#ns) = (Oc\<^bsup>n+1\<^esup>) @ [Bk] @ (tape_of_nat_list ns)" + +defs (overloaded) + tape_of_nl_abv: " \ tape_of_nat_list am" + tape_of_nat_abv : "<(n::nat)> \ Oc\<^bsup>n+1\<^esup>" + +text {* + @{text "crsp_l acf tcf"} meams the abacus configuration @{text "acf"} + is corretly represented by the TM configuration @{text "tcf"}. +*} + +fun crsp_l :: "layout \ abc_conf_l \ t_conf \ block list \ bool" + where + "crsp_l ly (as, lm) (ts, (l, r)) inres = + (ts = start_of ly as \ (\ rn. r = @ Bk\<^bsup>rn\<^esup>) + \ l = Bk # Bk # inres)" + +declare crsp_l.simps[simp del] + +subsection {* + A more general definition of TM execution. +*} + +(* +fun nnth_of :: "(taction \ nat) list \ nat \ nat \ (taction \ nat)" + where + "nnth_of p s b = (if 2*s < length p + then (p ! (2*s + b)) else (Nop, 0))" + +thm nth_of.simps + +fun nfetch :: "tprog \ nat \ block \ taction \ nat" + where + "nfetch p 0 b = (Nop, 0)" | + "nfetch p (Suc s) b = + (case b of + Bk \ nnth_of p s 0 | + Oc \ nnth_of p s 1)" +*) + +text {* + @{text "t_step tcf (tp, ss)"} returns the result of one step exection of TM @{text "tp"} + assuming @{text "tp"} starts from instial state @{text "ss"}. +*} + +fun t_step :: "t_conf \ (tprog \ nat) \ t_conf" + where + "t_step c (p, off) = + (let (state, leftn, rightn) = c in + let (action, next_state) = fetch p (state-off) + (case rightn of + [] \ Bk | + Bk # xs \ Bk | + Oc # xs \ Oc + ) + in + (next_state, new_tape action (leftn, rightn)))" + + +text {* + @{text "t_steps tcf (tp, ss) n"} returns the result of @{text "n"}-step exection + of TM @{text "tp"} assuming @{text "tp"} starts from instial state @{text "ss"}. +*} + +fun t_steps :: "t_conf \ (tprog \ nat) \ nat \ t_conf" + where + "t_steps c (p, off) 0 = c" | + "t_steps c (p, off) (Suc n) = t_steps + (t_step c (p, off)) (p, off) n" + +lemma stepn: "t_steps c (p, off) (Suc n) = + t_step (t_steps c (p, off) n) (p, off)" +apply(induct n arbitrary: c, simp add: t_steps.simps) +apply(simp add: t_steps.simps) +done + +text {* + The type of invarints expressing correspondence between + Abacus configuration and TM configuration. +*} + +type_synonym inc_inv_t = "abc_conf_l \ t_conf \ block list \ bool" + +declare tms_of.simps[simp del] tm_of.simps[simp del] + layout_of.simps[simp del] abc_fetch.simps [simp del] + t_step.simps[simp del] t_steps.simps[simp del] + tpairs_of.simps[simp del] start_of.simps[simp del] + fetch.simps [simp del] t_ncorrect.simps[simp del] + new_tape.simps [simp del] ci.simps [simp del] length_of.simps[simp del] + layout_of.simps[simp del] crsp_l.simps[simp del] + abc2t_correct.simps[simp del] + +lemma tct_div2: "t_ncorrect tp \ (length tp) mod 2 = 0" +apply(simp add: t_ncorrect.simps) +done + +lemma t_shift_fetch: + "\t_ncorrect tp1; t_ncorrect tp; + length tp1 div 2 < a \ a \ length tp1 div 2 + length tp div 2\ + \ fetch tp (a - length tp1 div 2) b = + fetch (tp1 @ tp @ tp2) a b" +apply(subgoal_tac "\ x. a = length tp1 div 2 + x", erule exE, simp) +apply(case_tac x, simp) +apply(subgoal_tac "length tp1 div 2 + Suc nat = + Suc (length tp1 div 2 + nat)") +apply(simp only: fetch.simps nth_of.simps, auto) +apply(case_tac b, simp) +apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp) +apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp) +apply(simp add: t_ncorrect.simps, auto) +apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp) +apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto) +apply(simp add: t_ncorrect.simps, auto) +apply(rule_tac x = "a - length tp1 div 2" in exI, simp) +done + +lemma t_shift_in_step: + "\t_step (a, aa, ba) (tp, length tp1 div 2) = (s, l, r); + t_ncorrect tp1; t_ncorrect tp; + length tp1 div 2 < a \ a \ length tp1 div 2 + length tp div 2\ + \ t_step (a, aa, ba) (tp1 @ tp @ tp2, 0) = (s, l, r)" +apply(simp add: t_step.simps) +apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \ + Bk | x # xs \ x) + = fetch (tp1 @ tp @ tp2) a (case ba of [] \ Bk | x # xs + \ x)") +apply(case_tac "fetch tp (a - length tp1 div 2) (case ba of [] \ Bk + | x # xs \ x)") +apply(auto intro: t_shift_fetch) +apply(case_tac ba, simp, simp) +apply(case_tac aaa, simp, simp) +done + +declare add_Suc_right[simp del] +lemma t_step_add: "t_steps c (p, off) (m + n) = + t_steps (t_steps c (p, off) m) (p, off) n" +apply(induct m arbitrary: n, simp add: t_steps.simps, simp) +apply(subgoal_tac "t_steps c (p, off) (Suc (m + n)) = + t_steps c (p, off) (m + Suc n)", simp) +apply(subgoal_tac "t_steps (t_steps c (p, off) m) (p, off) (Suc n) = + t_steps (t_step (t_steps c (p, off) m) (p, off)) + (p, off) n") +apply(simp, simp add: stepn) +apply(simp only: t_steps.simps) +apply(simp only: add_Suc_right) +done +declare add_Suc_right[simp] + +lemma s_out_fetch: "\t_ncorrect tp; + \ (length tp1 div 2 < a \ a \ length tp1 div 2 + + length tp div 2)\ + \ fetch tp (a - length tp1 div 2) b = (Nop, 0)" +apply(auto) +apply(simp add: fetch.simps) +apply(subgoal_tac "\ x. a - length tp1 div 2 = length tp div 2 + x") +apply(erule exE, simp) +apply(case_tac x, simp) +apply(auto simp add: fetch.simps) +apply(subgoal_tac "2 * (length tp div 2) = length tp") +apply(auto simp: t_ncorrect.simps split: block.splits) +apply(rule_tac x = "a - length tp1 div 2 - length tp div 2" in exI + , simp) +done + +lemma conf_keep_step: + "\t_ncorrect tp; + \ (length tp1 div 2 < a \ a \ length tp1 div 2 + + length tp div 2)\ + \ t_step (a, aa, ba) (tp, length tp1 div 2) = (0, aa, ba)" +apply(simp add: t_step.simps) +apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \ + Bk | Bk # xs \ Bk | Oc # xs \ Oc) = (Nop, 0)") +apply(simp add: new_tape.simps) +apply(rule s_out_fetch, simp, simp) +done + +lemma conf_keep: + "\t_ncorrect tp; + \ (length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2); n > 0\ + \ t_steps (a, aa, ba) (tp, length tp1 div 2) n = (0, aa, ba)" +apply(induct n, simp) +apply(case_tac n, simp add: t_steps.simps) +apply(rule_tac conf_keep_step, simp+) +apply(subgoal_tac " t_steps (a, aa, ba) + (tp, length tp1 div 2) (Suc (Suc nat)) + = t_step (t_steps (a, aa, ba) + (tp, length tp1 div 2) (Suc nat)) (tp, length tp1 div 2)") +apply(simp) +apply(rule_tac conf_keep_step, simp, simp) +apply(rule stepn) +done + +lemma state_bef_inside: + "\t_ncorrect tp1; t_ncorrect tp; + t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); + length tp1 div 2 < s0 \ + s0 \ length tp1 div 2 + length tp div 2; + length tp1 div 2 < s \ s \ length tp1 div 2 + length tp div 2; + n < stp; t_steps (s0, l0, r0) (tp, length tp1 div 2) n = + (a, aa, ba)\ + \ length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2" +apply(subgoal_tac "\ x. stp = n + x", erule exE) +apply(simp only: t_step_add) +apply(rule classical) +apply(subgoal_tac "t_steps (a, aa, ba) + (tp, length tp1 div 2) x = (0, aa, ba)") +apply(simp) +apply(rule conf_keep, simp, simp, simp) +apply(rule_tac x = "stp - n" in exI, simp) +done + +lemma turing_shift_inside: + "\t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); + length tp1 div 2 < s0 \ + s0 \ length tp1 div 2 + length tp div 2; + t_ncorrect tp1; t_ncorrect tp; + length tp1 div 2 < s \ + s \ length tp1 div 2 + length tp div 2\ + \ t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = (s, l, r)" +apply(induct stp arbitrary: s l r) +apply(simp add: t_steps.simps) +apply(subgoal_tac " t_steps (s0, l0, r0) + (tp, length tp1 div 2) (Suc stp) + = t_step (t_steps (s0, l0, r0) + (tp, length tp1 div 2) stp) (tp, length tp1 div 2)") +apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) stp") +apply(subgoal_tac "length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2") +apply(subgoal_tac "t_steps (s0, l0, r0) + (tp1 @ tp @ tp2, 0) stp = (a, b, c)") +apply(simp only: stepn, simp) +apply(rule_tac t_shift_in_step, simp+) +defer +apply(rule stepn) +apply(rule_tac n = stp and stp = "Suc stp" and a = a + and aa = b and ba = c in state_bef_inside, simp+) +done + +lemma take_Suc_last[elim]: "Suc as \ length xs \ + take (Suc as) xs = take as xs @ [xs ! as]" +apply(induct xs arbitrary: as, simp, simp) +apply(case_tac as, simp, simp) +done + +lemma concat_suc: "Suc as \ length xs \ + concat (take (Suc as) xs) = concat (take as xs) @ xs! as" +apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp) +by auto + +lemma concat_take_suc_iff: "Suc n \ length tps \ + concat (take n tps) @ (tps ! n) = concat (take (Suc n) tps)" +apply(drule_tac concat_suc, simp) +done + +lemma concat_drop_suc_iff: + "Suc n < length tps \ concat (drop (Suc n) tps) = + tps ! Suc n @ concat (drop (Suc (Suc n)) tps)" +apply(induct tps arbitrary: n, simp, simp) +apply(case_tac tps, simp, simp) +apply(case_tac n, simp, simp) +done + +declare append_assoc[simp del] + +lemma tm_append: "\n < length tps; tp = tps ! n\ \ + \ tp1 tp2. concat tps = tp1 @ tp @ tp2 \ tp1 = + concat (take n tps) \ tp2 = concat (drop (Suc n) tps)" +apply(rule_tac x = "concat (take n tps)" in exI) +apply(rule_tac x = "concat (drop (Suc n) tps)" in exI) +apply(auto) +apply(induct n, simp) +apply(case_tac tps, simp, simp, simp) +apply(subgoal_tac "concat (take n tps) @ (tps ! n) = + concat (take (Suc n) tps)") +apply(simp only: append_assoc[THEN sym], simp only: append_assoc) +apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @ + concat (drop (Suc (Suc n)) tps)", simp) +apply(rule_tac concat_drop_suc_iff, simp) +apply(rule_tac concat_take_suc_iff, simp) +done + +declare append_assoc[simp] + +lemma map_of: "n < length xs \ (map f xs) ! n = f (xs ! n)" +by(auto) + +lemma [simp]: "length (tms_of aprog) = length aprog" +apply(auto simp: tms_of.simps tpairs_of.simps) +done + +lemma ci_nth: "\ly = layout_of aprog; as < length aprog; + abc_fetch as aprog = Some ins\ + \ ci ly (start_of ly as) ins = tms_of aprog ! as" +apply(simp add: tms_of.simps tpairs_of.simps + abc_fetch.simps map_of del: map_append) +done + +lemma t_split:"\ + ly = layout_of aprog; + as < length aprog; abc_fetch as aprog = Some ins\ + \ \ tp1 tp2. concat (tms_of aprog) = + tp1 @ (ci ly (start_of ly as) ins) @ tp2 + \ tp1 = concat (take as (tms_of aprog)) \ + tp2 = concat (drop (Suc as) (tms_of aprog))" +apply(insert tm_append[of "as" "tms_of aprog" + "ci ly (start_of ly as) ins"], simp) +apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as") +apply(subgoal_tac "length (tms_of aprog) = length aprog", simp, simp) +apply(rule_tac ci_nth, auto) +done + +lemma math_sub: "\x >= Suc 0; x - 1 = z\ \ x + y - Suc 0 = z + y" +by auto + +lemma start_more_one: "as \ 0 \ start_of ly as >= Suc 0" +apply(induct as, simp add: start_of.simps) +apply(case_tac as, auto simp: start_of.simps) +done + +lemma tm_ct: "\abc2t_correct aprog; tp \ set (tms_of aprog)\ \ + t_ncorrect tp" +apply(simp add: abc2t_correct.simps tms_of.simps) +apply(auto) +apply(simp add:list_all_iff, auto) +done + +lemma div_apart: "\x mod (2::nat) = 0; y mod 2 = 0\ + \ (x + y) div 2 = x div 2 + y div 2" +apply(drule mod_eqD)+ +apply(auto) +done + +lemma div_apart_iff: "\x mod (2::nat) = 0; y mod 2 = 0\ \ + (x + y) mod 2 = 0" +apply(auto) +done + +lemma tms_ct: "\abc2t_correct aprog; n < length aprog\ \ + t_ncorrect (concat (take n (tms_of aprog)))" +apply(induct n, simp add: t_ncorrect.simps, simp) +apply(subgoal_tac "concat (take (Suc n) (tms_of aprog)) = + concat (take n (tms_of aprog)) @ (tms_of aprog ! n)", simp) +apply(simp add: t_ncorrect.simps) +apply(rule_tac div_apart_iff, simp) +apply(subgoal_tac "t_ncorrect (tms_of aprog ! n)", + simp add: t_ncorrect.simps) +apply(rule_tac tm_ct, simp) +apply(rule_tac nth_mem, simp add: tms_of.simps tpairs_of.simps) +apply(rule_tac concat_suc, simp add: tms_of.simps tpairs_of.simps) +done + +lemma tcorrect_div2: "\abc2t_correct aprog; Suc as < length aprog\ + \ (length (concat (take as (tms_of aprog))) + length (tms_of aprog + ! as)) div 2 = length (concat (take as (tms_of aprog))) div 2 + + length (tms_of aprog ! as) div 2" +apply(subgoal_tac "t_ncorrect (tms_of aprog ! as)") +apply(subgoal_tac "t_ncorrect (concat (take as (tms_of aprog)))") +apply(rule_tac div_apart) +apply(rule tct_div2, simp)+ +apply(erule_tac tms_ct, simp) +apply(rule_tac tm_ct, simp) +apply(rule_tac nth_mem) +apply(simp add: tms_of.simps tpairs_of.simps) +done + +lemma [simp]: "length (layout_of aprog) = length aprog" +apply(auto simp: layout_of.simps) +done + +lemma start_of_ind: "\as < length aprog; ly = layout_of aprog\ \ + start_of ly (Suc as) = start_of ly as + + length ((tms_of aprog) ! as) div 2" +apply(simp only: start_of.simps, simp) +apply(auto simp: start_of.simps tms_of.simps layout_of.simps + tpairs_of.simps) +apply(simp add: ci_length) +done + +lemma concat_take_suc: "Suc n \ length xs \ + concat (take (Suc n) xs) = concat (take n xs) @ (xs ! n)" +apply(subgoal_tac "take (Suc n) xs = + take n xs @ [xs ! n]") +apply(auto) +done + +lemma ci_length_not0: "Suc 0 <= length (ci ly as i) div 2" +apply(subgoal_tac "length (ci ly as i) div 2 = length_of i") +apply(simp add: length_of.simps split: abc_inst.splits) +apply(rule ci_length) +done + +lemma findnth_length2: "length (findnth n) = 4 * n" +apply(induct n, simp) +apply(simp) +done + +lemma ci_length2: "length (ci ly as i) = 2 * (length_of i)" +apply(simp add: ci.simps length_of.simps tinc_b_def tdec_b_def + split: abc_inst.splits, auto) +apply(simp add: findnth_length2)+ +done + +lemma tm_mod2: "as < length aprog \ + length (tms_of aprog ! as) mod 2 = 0" +apply(simp add: tms_of.simps) +apply(subgoal_tac "map (\(x, y). ci (layout_of aprog) x y) + (tpairs_of aprog) ! as + = (\(x, y). ci (layout_of aprog) x y) + ((tpairs_of aprog) ! as)", simp) +apply(case_tac "(tpairs_of aprog ! as)", simp) +apply(subgoal_tac "length (ci (layout_of aprog) a b) = + 2 * (length_of b)", simp) +apply(rule ci_length2) +apply(rule map_of, simp add: tms_of.simps tpairs_of.simps) +done + +lemma tms_mod2: "as \ length aprog \ + length (concat (take as (tms_of aprog))) mod 2 = 0" +apply(induct as, simp, simp) +apply(subgoal_tac "concat (take (Suc as) (tms_of aprog)) + = concat (take as (tms_of aprog)) @ + (tms_of aprog ! as)", auto) +apply(rule div_apart_iff, simp, rule tm_mod2, simp) +apply(rule concat_take_suc, simp add: tms_of.simps tpairs_of.simps) +done + +lemma [simp]: "\as < length aprog; (abc_fetch as aprog) = Some ins\ + \ ci (layout_of aprog) + (start_of (layout_of aprog) as) (ins) \ set (tms_of aprog)" +apply(insert ci_nth[of "layout_of aprog" aprog as], simp) +done + +lemma startof_not0: "start_of ly as > 0" +apply(induct as, simp add: start_of.simps) +apply(case_tac as, auto simp: start_of.simps) +done + +declare abc_step_l.simps[simp del] +lemma pre_lheq: "\tp = concat (take as (tms_of aprog)); + abc2t_correct aprog; as \ length aprog\ \ + start_of (layout_of aprog) as - Suc 0 = length tp div 2" +apply(induct as arbitrary: tp, simp add: start_of.simps, simp) +proof - + fix as tp + assume h1: "\tp. tp = concat (take as (tms_of aprog)) \ + start_of (layout_of aprog) as - Suc 0 = + length (concat (take as (tms_of aprog))) div 2" + and h2: " abc2t_correct aprog" "Suc as \ length aprog" + from h2 show "start_of (layout_of aprog) (Suc as) - Suc 0 = + length (concat (take (Suc as) (tms_of aprog))) div 2" + apply(insert h1[of "concat (take as (tms_of aprog))"], simp) + apply(insert start_of_ind[of as aprog "layout_of aprog"], simp) + apply(subgoal_tac "(take (Suc as) (tms_of aprog)) = + take as (tms_of aprog) @ [(tms_of aprog) ! as]", simp) + apply(subgoal_tac "(length (concat (take as (tms_of aprog))) + + length (tms_of aprog ! as)) div 2 + = length (concat (take as (tms_of aprog))) div 2 + + length (tms_of aprog ! as) div 2", simp) + apply(subgoal_tac "start_of (layout_of aprog) as = + length (concat (take as (tms_of aprog))) div 2 + Suc 0", simp) + apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp, + rule_tac startof_not0) + apply(insert tm_mod2[of as aprog], simp) + apply(insert tms_mod2[of as aprog], simp, arith) + apply(rule take_Suc_last, simp) + done +qed + +lemma crsp2stateq: + "\as < length aprog; abc2t_correct aprog; + crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\ \ + a = length (concat (take as (tms_of aprog))) div 2 + 1" +apply(simp add: crsp_l.simps) +apply(insert pre_lheq[of "(concat (take as (tms_of aprog)))" as aprog] +, simp) +apply(subgoal_tac "start_of (layout_of aprog) as > 0", + auto intro: startof_not0) +done + +lemma turing_shift_outside: + "\t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); + s \ 0; stp > 0; + length tp1 div 2 < s0 \ + s0 \ length tp1 div 2 + length tp div 2; + t_ncorrect tp1; t_ncorrect tp; + \ (length tp1 div 2 < s \ + s \ length tp1 div 2 + length tp div 2)\ + \ \stp' > 0. t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp' + = (s, l, r)" +apply(rule_tac x = stp in exI) +apply(case_tac stp, simp add: t_steps.simps) +apply(simp only: stepn) +apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) nat") +apply(subgoal_tac "length tp1 div 2 < a \ + a \ length tp1 div 2 + length tp div 2") +apply(subgoal_tac "t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) nat + = (a, b, c)", simp) +apply(rule_tac t_shift_in_step, simp+) +apply(rule_tac turing_shift_inside, simp+) +apply(rule classical) +apply(subgoal_tac "t_step (a,b,c) + (tp, length tp1 div 2) = (0, b, c)", simp) +apply(rule_tac conf_keep_step, simp+) +done + +lemma turing_shift: + "\t_steps (s0, (l0, r0)) (tp, (length tp1 div 2)) stp + = (s, (l, r)); s \ 0; stp > 0; + (length tp1 div 2 < s0 \ s0 <= length tp1 div 2 + length tp div 2); + t_ncorrect tp1; t_ncorrect tp\ \ + \ stp' > 0. t_steps (s0, (l0, r0)) (tp1 @ tp @ tp2, 0) stp' = + (s, (l, r))" +apply(case_tac "s > length tp1 div 2 \ + s <= length tp1 div 2 + length tp div 2") +apply(subgoal_tac " t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = + (s, l, r)") +apply(rule_tac x = stp in exI, simp) +apply(rule_tac turing_shift_inside, simp+) +apply(rule_tac turing_shift_outside, simp+) +done + +lemma inc_startof_not0: "start_of ly as \ Suc 0" +apply(induct as, simp add: start_of.simps) +apply(simp add: start_of.simps) +done + +lemma s_crsp: + "\as < length aprog; abc_fetch as aprog = Some ins; + abc2t_correct aprog; + crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\ \ + length (concat (take as (tms_of aprog))) div 2 < a + \ a \ length (concat (take as (tms_of aprog))) div 2 + + length (ci (layout_of aprog) (start_of (layout_of aprog) as) + ins) div 2" +apply(subgoal_tac "a = length (concat (take as (tms_of aprog))) div + 2 + 1", simp) +apply(rule_tac ci_length_not0) +apply(rule crsp2stateq, simp+) +done + +lemma tms_out_ex: + "\ly = layout_of aprog; tprog = tm_of aprog; + abc2t_correct aprog; + crsp_l ly (as, am) tc inres; as < length aprog; + abc_fetch as aprog = Some ins; + t_steps tc (ci ly (start_of ly as) ins, + (start_of ly as) - 1) n = (s, l, r); + n > 0; + abc_step_l (as, am) (abc_fetch as aprog) = (as', am'); + s = start_of ly as' + \ + \ \ stp > 0. (t_steps tc (tprog, 0) stp = (s, (l, r)))" +apply(simp only: tm_of.simps) +apply(subgoal_tac "\ tp1 tp2. concat (tms_of aprog) = + tp1 @ (ci ly (start_of ly as) ins) @ tp2 + \ tp1 = concat (take as (tms_of aprog)) \ + tp2 = concat (drop (Suc as) (tms_of aprog))") +apply(erule exE, erule exE, erule conjE, erule conjE, + case_tac tc, simp) +apply(rule turing_shift) +apply(subgoal_tac "start_of (layout_of aprog) as - Suc 0 + = length tp1 div 2", simp) +apply(rule_tac pre_lheq, simp, simp, simp) +apply(simp add: startof_not0, simp) +apply(rule_tac s_crsp, simp, simp, simp, simp) +apply(rule tms_ct, simp, simp) +apply(rule tm_ct, simp) +apply(subgoal_tac "ci (layout_of aprog) + (start_of (layout_of aprog) as) ins + = (tms_of aprog ! as)", simp) +apply(simp add: tms_of.simps tpairs_of.simps) +apply(simp add: tms_of.simps tpairs_of.simps abc_fetch.simps) +apply(erule_tac t_split, auto simp: tm_of.simps) +done + +(* +subsection {* The compilation of @{text "Inc n"} *} +*) + +text {* + The lemmas in this section lead to the correctness of + the compilation of @{text "Inc n"} instruction. +*} + +fun at_begin_fst_bwtn :: "inc_inv_t" + where + "at_begin_fst_bwtn (as, lm) (s, l, r) ires = + (\ lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \ length lm1 = s \ + (if lm1 = [] then l = Bk # Bk # ires + else l = [Bk]@@Bk#Bk#ires) \ r = (Bk\<^bsup>rn\<^esup>))" + + +fun at_begin_fst_awtn :: "inc_inv_t" + where + "at_begin_fst_awtn (as, lm) (s, l, r) ires = + (\ lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \ length lm1 = s \ + (if lm1 = [] then l = Bk # Bk # ires + else l = [Bk]@@Bk#Bk#ires) \ r = [Oc]@Bk\<^bsup>rn\<^esup> + )" + +fun at_begin_norm :: "inc_inv_t" + where + "at_begin_norm (as, lm) (s, l, r) ires= + (\ lm1 lm2 rn. lm = lm1 @ lm2 \ length lm1 = s \ + (if lm1 = [] then l = Bk # Bk # ires + else l = Bk # @ Bk# Bk # ires ) \ r = @ (Bk\<^bsup>rn\<^esup>))" + +fun in_middle :: "inc_inv_t" + where + "in_middle (as, lm) (s, l, r) ires = + (\ lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2 + \ length lm1 = s \ m + 1 = ml + mr \ + ml \ 0 \ tn = s + 1 - length lm \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>)@[Bk]@@ + Bk # Bk # ires) \ (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>) \ + (lm2 = [] \ r = (Oc\<^bsup>mr\<^esup>))) + )" + +fun inv_locate_a :: "inc_inv_t" + where "inv_locate_a (as, lm) (s, l, r) ires = + (at_begin_norm (as, lm) (s, l, r) ires \ + at_begin_fst_bwtn (as, lm) (s, l, r) ires \ + at_begin_fst_awtn (as, lm) (s, l, r) ires + )" + +fun inv_locate_b :: "inc_inv_t" + where "inv_locate_b (as, lm) (s, l, r) ires = + (in_middle (as, lm) (s, l, r)) ires " + +fun inv_after_write :: "inc_inv_t" + where "inv_after_write (as, lm) (s, l, r) ires = + (\ rn m lm1 lm2. lm = lm1 @ m # lm2 \ + (if lm1 = [] then l = Oc\<^bsup>m\<^esup> @ Bk # Bk # ires + else Oc # l = Oc\<^bsup>Suc m \<^esup>@ Bk # @ + Bk # Bk # ires) \ r = [Oc] @ @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_after_move :: "inc_inv_t" + where "inv_after_move (as, lm) (s, l, r) ires = + (\ rn m lm1 lm2. lm = lm1 @ m # lm2 \ + (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires + else l = Oc\<^bsup>Suc m\<^esup>@ Bk # @ Bk # Bk # ires) \ + r = @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_after_clear :: "inc_inv_t" + where "inv_after_clear (as, lm) (s, l, r) ires = + (\ rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \ + (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires + else l = Oc\<^bsup>Suc m\<^esup>@ Bk # @ Bk # Bk # ires) \ + r = Bk # r' \ Oc # r' = @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_on_right_moving :: "inc_inv_t" + where "inv_on_right_moving (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = m \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>)) \ + (r = (Oc\<^bsup>mr\<^esup>) \ lm2 = [])))" + +fun inv_on_left_moving_norm :: "inc_inv_t" + where "inv_on_left_moving_norm (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ mr > 0 \ (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ Bk # @ Bk # Bk # ires) + \ (r = (Oc\<^bsup>mr\<^esup>) @ Bk # @ (Bk\<^bsup>rn\<^esup>) \ + (lm2 = [] \ r = Oc\<^bsup>mr\<^esup>)))" + +fun inv_on_left_moving_in_middle_B:: "inc_inv_t" + where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires = + (\ lm1 lm2 rn. lm = lm1 @ lm2 \ + (if lm1 = [] then l = Bk # ires + else l = @ Bk # Bk # ires) \ + r = Bk # @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_on_left_moving :: "inc_inv_t" + where "inv_on_left_moving (as, lm) (s, l, r) ires = + (inv_on_left_moving_norm (as, lm) (s, l, r) ires \ + inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)" + + +fun inv_check_left_moving_on_leftmost :: "inc_inv_t" + where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires = + (\ rn. l = ires \ r = [Bk, Bk] @ @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_check_left_moving_in_middle :: "inc_inv_t" + where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires = + + (\ lm1 lm2 r' rn. lm = lm1 @ lm2 \ + (Oc # l = @ Bk # Bk # ires) \ r = Oc # Bk # r' \ + r' = @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_check_left_moving :: "inc_inv_t" + where "inv_check_left_moving (as, lm) (s, l, r) ires = + (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \ + inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)" + +fun inv_after_left_moving :: "inc_inv_t" + where "inv_after_left_moving (as, lm) (s, l, r) ires= + (\ rn. l = Bk # ires \ r = Bk # @ (Bk\<^bsup>rn\<^esup>))" + +fun inv_stop :: "inc_inv_t" + where "inv_stop (as, lm) (s, l, r) ires= + (\ rn. l = Bk # Bk # ires \ r = @ (Bk\<^bsup>rn\<^esup>))" + + +fun inc_inv :: "layout \ nat \ inc_inv_t" + where + "inc_inv ly n (as, lm) (s, l, r) ires = + (let ss = start_of ly as in + let lm' = abc_lm_s lm n ((abc_lm_v lm n)+1) in + if s = 0 then False + else if s < ss then False + else if s < ss + 2 * n then + if (s - ss) mod 2 = 0 then + inv_locate_a (as, lm) ((s - ss) div 2, l, r) ires + else inv_locate_b (as, lm) ((s - ss) div 2, l, r) ires + else if s = ss + 2 * n then + inv_locate_a (as, lm) (n, l, r) ires + else if s = ss + 2 * n + 1 then + inv_locate_b (as, lm) (n, l, r) ires + else if s = ss + 2 * n + 2 then + inv_after_write (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 3 then + inv_after_move (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 4 then + inv_after_clear (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 5 then + inv_on_right_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 6 then + inv_on_left_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 7 then + inv_check_left_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 8 then + inv_after_left_moving (as, lm') (s - ss, l, r) ires + else if s = ss + 2 * n + 9 then + inv_stop (as, lm') (s - ss, l, r) ires + else False) " + +lemma fetch_intro: + "\\xs.\ba = Oc # xs\ \ P (fetch prog i Oc); + \xs.\ba = Bk # xs\ \ P (fetch prog i Bk); + ba = [] \ P (fetch prog i Bk) + \ \ P (fetch prog i + (case ba of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))" +by (auto split:list.splits block.splits) + +lemma length_findnth[simp]: "length (findnth n) = 4 * n" +apply(induct n, simp) +apply(simp) +done + +declare tshift.simps[simp del] +declare findnth.simps[simp del] + +lemma findnth_nth: + "\n > q; x < 4\ \ + (findnth n) ! (4 * q + x) = (findnth (Suc q) ! (4 * q + x))" +apply(induct n, simp) +apply(case_tac "q < n", simp add: findnth.simps, auto) +apply(simp add: nth_append) +apply(subgoal_tac "q = n", simp) +apply(arith) +done + +lemma Suc_pre[simp]: "\ a < start_of ly as \ + (Suc a - start_of ly as) = Suc (a - start_of ly as)" +apply(arith) +done + +lemma fetch_locate_a_o: " +\a q xs. + \\ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + a - start_of (layout_of aprog) as = 2 * q; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (2 * q)) Oc) = (R, a+1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append Suc_pre) +apply(subgoal_tac "(findnth n ! Suc (4 * q)) = + findnth (Suc q) ! (4 * q + 1)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 1) = + findnth (Suc q) ! (4 * q + 1)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma fetch_locate_a_b: " +\a q xs. + \abc_fetch as aprog = Some (Inc n); + \ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + a - start_of (layout_of aprog) as = 2 * q; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * q)) Bk) + = (W1, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + tshift.simps nth_append) +apply(subgoal_tac "(findnth n ! (4 * q)) = + findnth (Suc q) ! (4 * q )") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 0) = + findnth (Suc q) ! (4 * q + 0)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma [intro]: "x mod 2 = Suc 0 \ \ q. x = Suc (2 * q)" +apply(drule mod_eqD, auto) +done + +lemma add3_Suc: "x + 3 = Suc (Suc (Suc x))" +apply(arith) +done + +declare start_of.simps[simp] +(* +lemma layout_not0: "start_of ly as > 0" +by(induct as, auto) +*) +lemma [simp]: + "\\ a < start_of (layout_of aprog) as; + a - start_of (layout_of aprog) as = Suc (2 * q); + abc_fetch as aprog = Some (Inc n); + start_of (layout_of aprog) as > 0\ + \ Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) = a" +apply(subgoal_tac +"Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) + = 2 + 2 * q + start_of (layout_of aprog) as - Suc 0", + simp, simp add: inc_startof_not0) +done + +lemma fetch_locate_b_o: " +\a xs. + \0 < a; \ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + (a - start_of (layout_of aprog) as) mod 2 = Suc 0; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Oc) = (R, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of (layout_of aprog) as) = + 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) + = findnth (Suc q) ! (4 * q + 3)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 3) = + findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc) +apply(rule_tac findnth_nth, auto) +done + +lemma fetch_locate_b_b: " +\a xs. + \0 < a; \ a < start_of (layout_of aprog) as; + a < start_of (layout_of aprog) as + 2 * n; + (a - start_of (layout_of aprog) as) mod 2 = Suc 0; + start_of (layout_of aprog) as > 0\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Bk) + = (R, a + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of (layout_of aprog) as) = + 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = + findnth (Suc q) ! (4 * q + 2)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 2) = + findnth (Suc q) ! (4 * q + 2)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma fetch_locate_n_a_o: + "start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Oc) = + (R, start_of (layout_of aprog) as + 2 * n + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_locate_n_a_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Bk) + = (W1, start_of (layout_of aprog) as + 2 * n)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_locate_n_b_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (Suc (2 * n))) Oc) = + (R, start_of (layout_of aprog) as + 2 * n + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_locate_n_b_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (Suc (2 * n))) Bk) = + (W1, start_of (layout_of aprog) as + 2 * n + 2)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_after_write_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n)) (Suc (Suc (Suc (2 * n)))) Oc) = + (R, start_of (layout_of aprog) as + 2*n + 3)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_after_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Oc) + = (W0, start_of (layout_of aprog) as + 2 * n + 4)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_after_move_b: " + start_of (layout_of aprog) as > 0 + \(fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Bk) + = (L, start_of (layout_of aprog) as + 2 * n + 6)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_clear_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (5 + 2 * n) Bk) + = (R, start_of (layout_of aprog) as + 2 * n + 5)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "5 + 2*n = Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_right_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Oc) + = (R, start_of (layout_of aprog) as + 2 * n + 5)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_right_move_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Bk) + = (W1, start_of (layout_of aprog) as + 2 * n + 2)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_left_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Oc) + = (L, start_of (layout_of aprog) as + 2 * n + 6)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tinc_b_def add3_Suc) +apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_left_move_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Bk) + = (L, start_of (layout_of aprog) as + 2 * n + 7)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_check_left_move_o: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Oc) + = (L, start_of (layout_of aprog) as + 2 * n + 6)" +apply(auto simp: ci.simps findnth.simps tshift.simps tinc_b_def) +apply(subgoal_tac "8 + 2 * n = Suc (2 * n + 7)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_check_left_move_b: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Bk) + = (R, start_of (layout_of aprog) as + 2 * n + 8) " +apply(auto simp: ci.simps findnth.simps + tshift.simps tinc_b_def add3_Suc) +apply(subgoal_tac "8 + 2*n= Suc (2*n + 7)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma fetch_after_left_move: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (9 + 2*n) Bk) + = (R, start_of (layout_of aprog) as + 2 * n + 9)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemma fetch_stop: " + start_of (layout_of aprog) as > 0 + \ (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) (10 + 2 *n) b) + = (Nop, 0)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def + split: block.splits) +done + +lemma fetch_state0: " + (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n)) 0 b) + = (Nop, 0)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tinc_b_def) +done + +lemmas fetch_simps = + fetch_locate_a_o fetch_locate_a_b fetch_locate_b_o fetch_locate_b_b + fetch_locate_n_a_b fetch_locate_n_a_o fetch_locate_n_b_o + fetch_locate_n_b_b fetch_after_write_o fetch_after_move_o + fetch_after_move_b fetch_clear_b fetch_right_move_o + fetch_right_move_b fetch_left_move_o fetch_left_move_b + fetch_after_left_move fetch_check_left_move_o fetch_stop + fetch_state0 fetch_check_left_move_b + +text {* *} +declare exponent_def[simp del] tape_of_nat_list.simps[simp del] + at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del] + at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del] + abc_lm_s.simps[simp del] abc_lm_v.simps[simp del] + ci.simps[simp del] t_step.simps[simp del] + inv_after_move.simps[simp del] + inv_on_left_moving_norm.simps[simp del] + inv_on_left_moving_in_middle_B.simps[simp del] + inv_after_clear.simps[simp del] + inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del] + inv_on_right_moving.simps[simp del] + inv_check_left_moving.simps[simp del] + inv_check_left_moving_in_middle.simps[simp del] + inv_check_left_moving_on_leftmost.simps[simp del] + inv_after_left_moving.simps[simp del] + inv_stop.simps[simp del] inv_locate_a.simps[simp del] + inv_locate_b.simps[simp del] +declare tms_of.simps[simp del] tm_of.simps[simp del] + layout_of.simps[simp del] abc_fetch.simps [simp del] + t_step.simps[simp del] t_steps.simps[simp del] + tpairs_of.simps[simp del] start_of.simps[simp del] + fetch.simps [simp del] new_tape.simps [simp del] + nth_of.simps [simp del] ci.simps [simp del] + length_of.simps[simp del] + +(*! Start point *) +lemma [simp]: "Suc (2 * q) mod 2 = Suc 0" +by arith + +lemma [simp]: "Suc (2 * q) div 2 = q" +by arith + +lemma [simp]: "\ \ a < start_of ly as; + a < start_of ly as + 2 * n; a - start_of ly as = 2 * q\ + \ Suc a < start_of ly as + 2 * n" +apply(arith) +done + +lemma [simp]: "x mod 2 = Suc 0 \ (Suc x) mod 2 = 0" +by arith + +lemma [simp]: "x mod 2 = Suc 0 \ (Suc x) div 2 = Suc (x div 2)" +by arith +lemma exp_def[simp]: "a\<^bsup>Suc n \<^esup>= a # a\<^bsup>n\<^esup>" +by(simp add: exponent_def) +lemma [intro]: "Bk # r = Oc\<^bsup>mr\<^esup> @ r' \ mr = 0" +by(case_tac mr, auto simp: exponent_def) + +lemma [intro]: "Bk # r = replicate mr Oc \ mr = 0" +by(case_tac mr, auto) +lemma tape_of_nl_abv_cons[simp]: "xs \ [] \ + = Oc\<^bsup>Suc x\<^esup>@ Bk # " +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac xs, simp, simp add: tape_of_nat_list.simps) +done + +lemma [simp]: "<[]::nat list> = []" +by(auto simp: tape_of_nl_abv tape_of_nat_list.simps) +lemma [simp]: "Oc # r = <(lm::nat list)> @ Bk\<^bsup>rn\<^esup>\ lm \ []" +apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, auto simp: exponent_def) +done +lemma BkCons_nil: "Bk # xs = @ Bk\<^bsup>rn\<^esup>\ lm = []" +apply(case_tac lm, simp) +apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done +lemma BkCons_nil': "Bk # xs = @ Bk\<^bsup>ln\<^esup>\ lm = []" +by(auto intro: BkCons_nil) + +lemma hd_tl_tape_of_nat_list: + "tl (lm::nat list) \ [] \ = @ Bk # " +apply(frule tape_of_nl_abv_cons[of "tl lm" "hd lm"]) +apply(simp add: tape_of_nat_abv Bk_def del: tape_of_nl_abv_cons) +apply(subgoal_tac "lm = hd lm # tl lm", auto) +apply(case_tac lm, auto) +done +lemma [simp]: "Oc # xs = Oc\<^bsup>mr\<^esup> @ Bk # @ Bk\<^bsup>rn\<^esup>\ mr > 0" +apply(case_tac mr, auto simp: exponent_def) +done + +lemma tape_of_nat_list_cons: "xs \ [] \ tape_of_nat_list (x # xs) = + replicate (Suc x) Oc @ Bk # tape_of_nat_list xs" +apply(drule tape_of_nl_abv_cons[of xs x]) +apply(auto simp: tape_of_nl_abv tape_of_nat_abv Oc_def Bk_def exponent_def) +done + +lemma rev_eq: "rev xs = rev ys \ xs = ys" +by simp + +lemma tape_of_nat_list_eq: " xs = ys \ + tape_of_nat_list xs = tape_of_nat_list ys" +by simp + +lemma tape_of_nl_nil_eq: "<(lm::nat list)> = [] = (lm = [])" +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac lm, simp add: tape_of_nat_list.simps) +apply(case_tac "list") +apply(auto simp: tape_of_nat_list.simps) +done + +lemma rep_ind: "replicate (Suc n) a = replicate n a @ [a]" +apply(induct n, simp, simp) +done + +lemma [simp]: "Oc # r = @ replicate rn Bk \ Suc 0 \ length lm" +apply(rule_tac classical, auto) +apply(case_tac lm, simp, case_tac rn, auto) +done +lemma Oc_Bk_Cons: "Oc # Bk # list = @ Bk\<^bsup>ln\<^esup> \ + lm \ [] \ hd lm = 0" +apply(case_tac lm, simp, case_tac ln, simp add: exponent_def, simp add: exponent_def, simp) +apply(case_tac lista, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done +(*lemma Oc_Oc_Cons: "Oc # Oc # list = @ Bk\<^bsup>ln\<^esup> \ + lm \ [] \ hd lm > 0" +apply(case_tac lm, simp add: exponent_def, case_tac ln, simp, simp) +apply(case_tac lista, + auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(case_tac [!] a, auto) +apply(case_tac ln, auto) +done +*) +lemma Oc_nil_zero[simp]: "[Oc] = @ Bk\<^bsup>ln\<^esup> + \ lm = [0] \ ln = 0" +apply(case_tac lm, simp) +apply(case_tac ln, auto simp: exponent_def) +apply(case_tac [!] list, + auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "Oc # r = @ replicate rn Bk \ + (\rn. r = replicate (hd lm2) Oc @ Bk # @ + replicate rn Bk) \ + tl lm2 = [] \ r = replicate (hd lm2) Oc" +apply(rule_tac disjCI, simp) +apply(case_tac "tl lm2 = []", simp) +apply(case_tac lm2, simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, simp, simp, simp) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(case_tac rn, simp, simp) +apply(rule_tac x = rn in exI) +apply(simp add: hd_tl_tape_of_nat_list) +apply(simp add: tape_of_nat_abv Oc_def exponent_def) +done + +(*inv: from locate_a to locate_b*) +lemma [simp]: + "inv_locate_a (as, lm) (q, l, Oc # r) ires + \ inv_locate_b (as, lm) (q, Oc # l, r) ires" +apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps + at_begin_norm.simps at_begin_fst_bwtn.simps + at_begin_fst_awtn.simps) +apply(erule disjE, erule exE, erule exE, erule exE) +apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp) +apply(rule_tac x = "0" in exI, rule_tac x = "hd lm2" in exI, + auto simp: exponent_def) +apply(rule_tac x = "Suc 0" in exI, simp add:exponent_def) +apply(rule_tac x = "lm @ replicate tn 0" in exI, + rule_tac x = "[]" in exI, + rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI) +apply(simp only: rep_ind, simp) +apply(rule_tac x = "Suc 0" in exI, auto) +apply(case_tac [1-3] rn, simp_all ) +apply(rule_tac x = "lm @ replicate tn 0" in exI, + rule_tac x = "[]" in exI, + rule_tac x = "Suc tn" in exI, + rule_tac x = 0 in exI, simp add: rep_ind del: replicate_Suc split:if_splits) +apply(rule_tac x = "Suc 0" in exI, auto) +apply(case_tac rn, simp, simp) +apply(rule_tac [!] x = "Suc 0" in exI, auto) +apply(case_tac [!] rn, simp_all) +done + +(*inv: from locate_a to _locate_a*) +lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires + \ inv_locate_a (as, am) (q, aaa, Oc # xs) ires" +apply(simp only: inv_locate_a.simps at_begin_norm.simps + at_begin_fst_bwtn.simps at_begin_fst_awtn.simps) +apply(erule_tac disjE, erule exE, erule exE, erule exE, + rule disjI2, rule disjI2) +defer +apply(erule_tac disjE, erule exE, erule exE, + erule exE, rule disjI2, rule disjI2) +prefer 2 +apply(simp) +proof- + fix lm1 tn rn + assume k: "lm1 = am @ 0\<^bsup>tn\<^esup> \ length lm1 = q \ (if lm1 = [] then aaa = Bk # Bk # + ires else aaa = [Bk] @ @ Bk # Bk # ires) \ Bk # xs = Bk\<^bsup>rn\<^esup>" + thus "\lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \ length lm1 = q \ (if lm1 = [] then + aaa = Bk # Bk # ires else aaa = [Bk] @ @ Bk # Bk # ires) \ Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>" + (is "\lm1 tn rn. ?P lm1 tn rn") + proof - + from k have "?P lm1 tn (rn - 1)" + apply(auto simp: Oc_def) + by(case_tac [!] "rn::nat", auto simp: exponent_def) + thus ?thesis by blast + qed +next + fix lm1 lm2 rn + assume h1: "am = lm1 @ lm2 \ length lm1 = q \ (if lm1 = [] + then aaa = Bk # Bk # ires else aaa = Bk # @ Bk # Bk # ires) \ + Bk # xs = @ Bk\<^bsup>rn\<^esup>" + from h1 have h2: "lm2 = []" + proof(rule_tac xs = xs and rn = rn in BkCons_nil, simp) + qed + from h1 and h2 show "\lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \ length lm1 = q \ + (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ @ Bk # Bk # ires) \ + Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>" + (is "\lm1 tn rn. ?P lm1 tn rn") + proof - + from h1 and h2 have "?P lm1 0 (rn - 1)" + apply(auto simp: Oc_def exponent_def + tape_of_nl_abv tape_of_nat_list.simps) + by(case_tac "rn::nat", simp, simp) + thus ?thesis by blast + qed +qed + +lemma [intro]: "\rn. [a] = a\<^bsup>rn\<^esup>" +by(rule_tac x = "Suc 0" in exI, simp add: exponent_def) + +lemma [intro]: "\tn. [] = a\<^bsup>tn\<^esup>" +apply(rule_tac x = 0 in exI, simp add: exponent_def) +done + +lemma [intro]: "at_begin_norm (as, am) (q, aaa, []) ires + \ at_begin_norm (as, am) (q, aaa, [Bk]) ires" +apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE) +apply(rule_tac x = lm1 in exI, simp, auto) +done + +lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires + \ at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires" +apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE) +apply(rule_tac x = "am @ 0\<^bsup>tn\<^esup>" in exI, auto) +done + +lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires + \ at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires" +apply(auto simp: at_begin_fst_awtn.simps) +done + +lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires + \ inv_locate_a (as, am) (q, aaa, [Bk]) ires" +apply(simp only: inv_locate_a.simps) +apply(erule disj_forward) +defer +apply(erule disj_forward, auto) +done + +lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \ + inv_locate_a (as, am) (q, aaa, [Oc]) ires" +apply(insert locate_a_2_locate_a [of as am q aaa "[]"]) +apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto) +done + +(*inv: from locate_b to locate_b*) +lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires + \ inv_locate_b (as, am) (q, Oc # aaa, xs) ires" +apply(simp only: inv_locate_b.simps in_middle.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = tn in exI, rule_tac x = m in exI) +apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI, + rule_tac x = rn in exI) +apply(case_tac mr, simp_all add: exponent_def, auto) +done +lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # @ + Bk\<^bsup>rn \<^esup>) \ (lm2 = [] \ Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>) + \ mr = 0 \ lm = []" +apply(rule context_conjI) +apply(case_tac mr, auto simp:exponent_def) +apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn]) +apply(case_tac n, auto simp: exponent_def Bk_def tape_of_nl_nil_eq) +done + +lemma tape_of_nat_def: "<[m::nat]> = Oc # Oc\<^bsup>m\<^esup>" +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) +done +lemma [simp]: "\inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \n. xs = Bk\<^bsup>n\<^esup>\ + \ inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" +apply(simp add: inv_locate_b.simps inv_locate_a.simps) +apply(rule_tac disjI2, rule_tac disjI1) +apply(simp only: in_middle.simps at_begin_fst_bwtn.simps) +apply(erule_tac exE)+ +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp) +apply(subgoal_tac "mr = 0 \ lm2 = []") +defer +apply(rule_tac n = n and mr = mr and lm = "lm2" + and rn = rn and n = n in zero_and_nil) +apply(auto simp: exponent_def) +apply(case_tac "lm1 = []", auto simp: tape_of_nat_def) +done + +lemma length_equal: "xs = ys \ length xs = length ys" +by auto +lemma [simp]: "a\<^bsup>0\<^esup> = []" +by(simp add: exp_zero) +(*inv: from locate_b to locate_a*) +lemma [simp]: "length (a\<^bsup>b\<^esup>) = b" +apply(simp add: exponent_def) +done + +lemma [simp]: "\inv_locate_b (as, am) (q, aaa, Bk # xs) ires; + \ (\n. xs = Bk\<^bsup>n\<^esup>)\ + \ inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" +apply(simp add: inv_locate_b.simps inv_locate_a.simps) +apply(rule_tac disjI1) +apply(simp only: in_middle.simps at_begin_norm.simps) +apply(erule_tac exE)+ +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp) +apply(subgoal_tac "tn = 0", simp add: exponent_def , auto split: if_splits) +apply(case_tac [!] mr, simp_all add: tape_of_nat_def, auto) +apply(case_tac lm2, simp, erule_tac x = rn in allE, simp) +apply(case_tac am, simp, simp) +apply(case_tac lm2, simp, erule_tac x = rn in allE, simp) +apply(drule_tac length_equal, simp) +done + +lemma locate_b_2_a[intro]: + "inv_locate_b (as, am) (q, aaa, Bk # xs) ires + \ inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" +apply(case_tac "\ n. xs = Bk\<^bsup>n\<^esup>", simp, simp) +done + +lemma locate_b_2_locate_a[simp]: + "\\ a < start_of ly as; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\ + \ (Suc a < start_of ly as + 2 * n \ inv_locate_a (as, am) + (Suc ((a - start_of ly as) div 2), Bk # aaa, xs) ires) \ + (\ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, xs) ires)" +apply(auto) +apply(subgoal_tac "n > 0") +apply(subgoal_tac "(a - start_of ly as) div 2 = n - 1") +apply(insert locate_b_2_a [of as am "n - 1" aaa xs], simp) +apply(arith) +apply(case_tac n, simp, simp) +done + +lemma [simp]: "inv_locate_b (as, am) (q, l, []) ires + \ inv_locate_b (as, am) (q, l, [Bk]) ires" +apply(simp only: inv_locate_b.simps in_middle.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = tn in exI, rule_tac x = m in exI, + rule_tac x = ml in exI, rule_tac x = mr in exI) +apply(auto) +done + +lemma locate_b_2_locate_a_B[simp]: + "\\ a < start_of ly as; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\ + \ (Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) + (Suc ((a - start_of ly as) div 2), Bk # aaa, []) ires) + \ (\ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, []) ires)" +apply(insert locate_b_2_locate_a [of a ly as n am aaa "[]"], simp) +done + +(*inv: from locate_b to after_write*) +lemma inv_locate_b_2_after_write[simp]: + "inv_locate_b (as, am) (n, aaa, Bk # xs) ires + \ inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) + (Suc (Suc (2 * n)), aaa, Oc # xs) ires" +apply(auto simp: in_middle.simps inv_after_write.simps + abc_lm_v.simps abc_lm_s.simps inv_locate_b.simps) +apply(subgoal_tac [!] "mr = 0", auto simp: exponent_def split: if_splits) +apply(subgoal_tac "lm2 = []", simp) +apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI, + rule_tac x = "lm1" in exI, simp, rule_tac x = "[]" in exI, simp) +apply(case_tac "Suc (length lm1) - length am", simp, simp only: rep_ind, simp) +apply(subgoal_tac "length lm1 - length am = nat", simp, arith) +apply(drule_tac length_equal, simp) +done + +lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \ + inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) + (Suc (Suc (2 * n)), aaa, [Oc]) ires" +apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"]) +by(simp) + +(*inv: from after_write to after_move*) +lemma [simp]: "inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires + \ inv_after_move (as, lm) (2 * n + 3, Oc # l, r) ires" +apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits) +done + +lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n) + )) (Suc (Suc (2 * n)), aaa, Bk # xs) ires = False" +apply(simp add: inv_after_write.simps ) +done + +lemma [simp]: + "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) + (Suc (Suc (2 * n)), aaa, []) ires = False" +apply(simp add: inv_after_write.simps ) +done + +(*inv: from after_move to after_clear*) +lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires + \ inv_after_clear (as, lm) (s', l, Bk # r) ires" +apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits) +done + +(*inv: from after_move to on_leftmoving*) +lemma inv_after_move_2_inv_on_left_moving[simp]: + "inv_after_move (as, lm) (s, l, Bk # r) ires + \ (l = [] \ + inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \ + (l \ [] \ + inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" +apply(simp only: inv_after_move.simps inv_on_left_moving.simps) +apply(subgoal_tac "l \ []", rule conjI, simp, rule impI, + rule disjI1, simp only: inv_on_left_moving_norm.simps) +apply(erule exE)+ +apply(subgoal_tac "lm2 = []") +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = m in exI, + rule_tac x = 1 in exI, + rule_tac x = "rn - 1" in exI, simp, case_tac rn) +apply(auto simp: exponent_def intro: BkCons_nil split: if_splits) +done + +lemma [elim]: "[] = \ lm = []" +using tape_of_nl_nil_eq[of lm] +by simp + +lemma inv_after_move_2_inv_on_left_moving_B[simp]: + "inv_after_move (as, lm) (s, l, []) ires + \ (l = [] \ inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \ + (l \ [] \ inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)" +apply(simp only: inv_after_move.simps inv_on_left_moving.simps) +apply(subgoal_tac "l \ []", rule conjI, simp, rule impI, rule disjI1, + simp only: inv_on_left_moving_norm.simps) +apply(erule exE)+ +apply(subgoal_tac "lm2 = []") +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = m in exI, + rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn) +apply(auto simp: exponent_def tape_of_nl_nil_eq intro: BkCons_nil split: if_splits) +done + +(*inv: from after_clear to on_right_moving*) +lemma [simp]: "Oc # r = replicate rn Bk = False" +apply(case_tac rn, simp, simp) +done + +lemma inv_after_clear_2_inv_on_right_moving[simp]: + "inv_after_clear (as, lm) (2 * n + 4, l, Bk # r) ires + \ inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, r) ires" +apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps ) +apply(subgoal_tac "lm2 \ []") +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, + rule_tac x = "hd lm2" in exI, simp) +apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI) +apply(simp add: exponent_def, rule conjI) +apply(case_tac [!] "lm2::nat list", auto simp: exponent_def) +apply(case_tac rn, auto split: if_splits simp: tape_of_nat_def) +apply(case_tac list, + simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(erule_tac x = "rn - 1" in allE, + case_tac rn, auto simp: exponent_def) +apply(case_tac list, + simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(erule_tac x = "rn - 1" in allE, + case_tac rn, auto simp: exponent_def) +done + + +lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires\ + inv_after_clear (as, lm) (2 * n + 4, l, [Bk]) ires" +by(auto simp: inv_after_clear.simps) + +lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires + \ inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, []) ires" +by(insert + inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp) + +(*inv: from on_right_moving to on_right_movign*) +lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, Oc # r) ires + \ inv_on_right_moving (as, lm) (2 * n + 5, Oc # l, r) ires" +apply(auto simp: inv_on_right_moving.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = "ml + mr" in exI, simp) +apply(rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, simp) +apply(case_tac mr, auto simp: exponent_def ) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = "ml + mr" in exI, simp) +apply(rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, simp) +apply(case_tac mr, auto split: if_splits simp: exponent_def) +done + +lemma inv_on_right_moving_2_inv_on_right_moving[simp]: + "inv_on_right_moving (as, lm) (2 * n + 5, l, Bk # r) ires + \ inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires" +apply(auto simp: inv_on_right_moving.simps inv_after_write.simps ) +apply(case_tac mr, auto simp: exponent_def split: if_splits) +apply(case_tac [!] mr, simp_all) +done + +lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires\ + inv_on_right_moving (as, lm) (2 * n + 5, l, [Bk]) ires" +apply(auto simp: inv_on_right_moving.simps exponent_def) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp) +apply (rule_tac x = m in exI, auto split: if_splits simp: exponent_def) +done + +(*inv: from on_right_moving to after_write*) +lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires + \ inv_after_write (as, lm) (Suc (Suc (2 * n)), l, [Oc]) ires" +apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp) +done + +(*inv: from on_left_moving to on_left_moving*) +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) + (s, l, Oc # r) ires = False" +apply(auto simp: inv_on_left_moving_in_middle_B.simps ) +done + +lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires + = False" +apply(auto simp: inv_on_left_moving_norm.simps) +apply(case_tac [!] mr, auto simp: ) +done + +lemma [intro]: "\rna. Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>rn\<^esup> = @ Bk\<^bsup>rna\<^esup>" +apply(case_tac lm, simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(rule_tac x = "Suc rn" in exI, simp) +apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto) +done + + +lemma [simp]: + "\inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; + hd l = Bk; l \ []\ \ + inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires" +apply(case_tac l, simp, simp) +apply(simp only: inv_on_left_moving_norm.simps + inv_on_left_moving_in_middle_B.simps) +apply(erule_tac exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto) +apply(case_tac [!] ml, auto) +apply(rule_tac [!] x = 0 in exI, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "\inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; + hd l = Oc; l \ []\ + \ inv_on_left_moving_norm (as, lm) + (s, tl l, Oc # Oc # r) ires" +apply(simp only: inv_on_left_moving_norm.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = "ml - 1" in exI, + rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp) +apply(case_tac ml, auto simp: exponent_def split: if_splits) +done + +lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires + \ inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires" +apply(auto simp: inv_on_left_moving_norm.simps + inv_on_left_moving_in_middle_B.simps split: if_splits) +done + +lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires + \ (l = [] \ inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires) + \ (l \ [] \ inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)" +apply(simp add: inv_on_left_moving.simps) +apply(case_tac "l \ []", rule conjI, simp, simp) +apply(case_tac "hd l", simp, simp, simp) +done + +(*inv: from on_left_moving to check_left_moving*) +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) + (s, Bk # list, Bk # r) ires + \ inv_check_left_moving_on_leftmost (as, lm) + (s', list, Bk # Bk # r) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps + inv_check_left_moving_on_leftmost.simps split: if_splits) +apply(case_tac [!] "rev lm1", simp_all) +apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: + "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False" +by(auto simp: inv_check_left_moving_in_middle.simps ) + +lemma [simp]: + "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\ + inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps + inv_check_left_moving_on_leftmost.simps split: if_splits) +done + + +lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm) + (s, list, Oc # r) ires= False" +by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits) + +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) + (s, Oc # list, Bk # r) ires + \ inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps + inv_check_left_moving_in_middle.simps split: if_splits) +done + +lemma inv_on_left_moving_2_check_left_moving[simp]: + "inv_on_left_moving (as, lm) (s, l, Bk # r) ires + \ (l = [] \ inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires) + \ (l \ [] \ + inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" +apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps) +apply(case_tac l, simp, simp) +apply(case_tac a, simp, simp) +done + +lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False" +apply(auto simp: inv_on_left_moving_norm.simps) +by(case_tac [!] mr, auto) + +lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\ + inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires" +apply(simp add: inv_on_left_moving.simps) +apply(auto simp: inv_on_left_moving_in_middle_B.simps) +done + +lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False" +apply(simp add: inv_on_left_moving.simps) +apply(simp add: inv_on_left_moving_in_middle_B.simps) +done + +lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires + \ (l = [] \ inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \ + (l \ [] \ inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)" +by simp + +lemma Oc_Bk_Cons_ex[simp]: + "Oc # Bk # list = @ Bk\<^bsup>ln\<^esup> \ + \ln. list = @ Bk\<^bsup>ln\<^esup>" +apply(case_tac "lm", simp) +apply(case_tac ln, simp_all add: exponent_def) +apply(case_tac lista, + auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def) +apply(case_tac [!] a, auto simp: ) +apply(case_tac ln, simp, rule_tac x = nat in exI, simp) +done + +lemma [simp]: + "Oc # Bk # list = @ Bk\<^bsup>ln\<^esup> \ + \rna. Oc # Bk # @ Bk\<^bsup>rn\<^esup> = @ Bk\<^bsup>rna\<^esup>" +apply(frule Oc_Bk_Cons, simp) +apply(case_tac lm2, + auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def ) +apply(rule_tac x = "Suc rn" in exI, simp) +done + +(*inv: from check_left_moving to on_left_moving*) +lemma [intro]: "\rna. a # a\<^bsup>rn\<^esup> = a\<^bsup>rna\<^esup>" +apply(rule_tac x = "Suc rn" in exI, simp) +done + +lemma +inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]: +"inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires + \ inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires" +apply(simp only: inv_check_left_moving_in_middle.simps + inv_on_left_moving_in_middle_B.simps) +apply(erule_tac exE)+ +apply(rule_tac x = "rev (tl (rev lm1))" in exI, + rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto) +apply(case_tac [!] "rev lm1",simp_all add: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac [!] a, simp_all) +apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps, auto) +apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps, auto) +apply(case_tac [!] lista, simp_all add: tape_of_nat_list.simps) +done + +lemma [simp]: + "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\ + inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires" +apply(auto simp: inv_check_left_moving_in_middle.simps ) +done + +lemma [simp]: + "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires + \ inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires" +apply(insert +inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of + as lm n "[]" r], simp) +done + +lemma [simp]: "a\<^bsup>0\<^esup> = []" +apply(simp add: exponent_def) +done + +lemma [simp]: "inv_check_left_moving_in_middle (as, lm) + (s, Oc # list, Oc # r) ires + \ inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires" +apply(auto simp: inv_check_left_moving_in_middle.simps + inv_on_left_moving_norm.simps) +apply(rule_tac x = "rev (tl (rev lm1))" in exI, + rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI) +apply(rule_tac conjI) +apply(case_tac "rev lm1", simp, simp) +apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto) +apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp) +apply(case_tac [!] "rev lm1", simp_all) +apply(case_tac [!] a, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto) +done + +lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires +\ (l = [] \ inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \ + (l \ [] \ inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)" +apply(case_tac l, + auto simp: inv_check_left_moving.simps inv_on_left_moving.simps) +apply(case_tac a, simp, simp) +done + +(*inv: check_left_moving to after_left_moving*) +lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires + \ inv_after_left_moving (as, lm) (s', Bk # l, r) ires" +apply(auto simp: inv_check_left_moving.simps + inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps) +done + + +lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires + \ inv_after_left_moving (as, lm) (s', Bk # l, []) ires" +by(simp add: inv_check_left_moving.simps +inv_check_left_moving_in_middle.simps +inv_check_left_moving_on_leftmost.simps) + +(*inv: after_left_moving to inv_stop*) +lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires + \ inv_stop (as, lm) (s', Bk # l, r) ires" +apply(auto simp: inv_after_left_moving.simps inv_stop.simps) +done + +lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires + \ inv_stop (as, lm) (s', Bk # l, []) ires" +by(auto simp: inv_after_left_moving.simps) + +(*inv: stop to stop*) +lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \ + inv_stop (as, lm) (y, l, r) ires" +apply(simp add: inv_stop.simps) +done + +lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False" +apply(auto simp: inv_after_clear.simps ) +done + +lemma [simp]: + "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False" +by(auto simp: inv_after_left_moving.simps ) + +lemma start_of_not0: "as \ 0 \ start_of ly as > 0" +apply(rule startof_not0) +done + +text {* + The single step currectness of the TM complied from Abacus instruction @{text "Inc n"}. + It shows every single step execution of this TM keeps the invariant. +*} + +lemma inc_inv_step: + assumes + -- {* Invariant holds on the start *} + h11: "inc_inv ly n (as, am) tc ires" + -- {* The layout of Abacus program @{text "aprog"} is @{text "ly"} *} + and h12: "ly = layout_of aprog" + -- {* The instruction at position @{text "as"} is @{text "Inc n"} *} + and h21: "abc_fetch as aprog = Some (Inc n)" + -- {* TM not yet reach the final state, where @{text "start_of ly as + 2*n + 9"} is the state + where the current TM stops and the next TM starts. *} + and h22: "(\ (s, l, r). s \ start_of ly as + 2*n + 9) tc" + shows + -- {* + Single step execution of the TM keeps the invaraint, where + the TM compiled from @{text "Inc n"} is @{text "(ci ly (start_of ly as) (Inc n))"} + @{text "start_of ly as - Suc 0)"} is the offset used to execute this {\em shifted} + TM. + *} + "inc_inv ly n (as, am) (t_step tc (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0)) ires" +proof - + from h21 h22 have h3 : "start_of (layout_of aprog) as > 0" + apply(case_tac as, simp add: start_of.simps abc_fetch.simps) + apply(insert start_of_not0[of as "layout_of aprog"], simp) + done + from h11 h12 and h21 h22 and this show ?thesis + apply(case_tac tc, simp) + apply(case_tac "a = 0", + auto split:if_splits simp add:t_step.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) + apply (simp_all add:fetch_simps new_tape.simps) + done +qed + + +lemma t_steps_ind: "t_steps tc (p, off) (Suc n) + = t_step (t_steps tc (p, off) n) (p, off)" +apply(induct n arbitrary: tc) +apply(simp add: t_steps.simps) +apply(simp add: t_steps.simps) +done + +definition lex_pair :: "((nat \ nat) \ (nat \ nat)) set" + where + "lex_pair \ less_than <*lex*> less_than" + +definition lex_triple :: + "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" + where "lex_triple \ less_than <*lex*> lex_pair" + +definition lex_square :: + "((nat \ nat \ nat \ nat) \ (nat \ nat \ nat \ nat)) set" + where "lex_square \ less_than <*lex*> lex_triple" + +fun abc_inc_stage1 :: "t_conf \ nat \ nat \ nat" + where + "abc_inc_stage1 (s, l, r) ss n = + (if s = 0 then 0 + else if s \ ss+2*n+1 then 5 + else if s\ ss+2*n+5 then 4 + else if s \ ss+2*n+7 then 3 + else if s = ss+2*n+8 then 2 + else 1)" + +fun abc_inc_stage2 :: "t_conf \ nat \ nat \ nat" + where + "abc_inc_stage2 (s, l, r) ss n = + (if s \ ss + 2*n + 1 then 0 + else if s = ss + 2*n + 2 then length r + else if s = ss + 2*n + 3 then length r + else if s = ss + 2*n + 4 then length r + else if s = ss + 2*n + 5 then + if r \ [] then length r + else 1 + else if s = ss+2*n+6 then length l + else if s = ss+2*n+7 then length l + else 0)" + +fun abc_inc_stage3 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_inc_stage3 (s, l, r) ss n ires = ( + if s = ss + 2*n + 3 then 4 + else if s = ss + 2*n + 4 then 3 + else if s = ss + 2*n + 5 then + if r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2*n + 2 then 0 + else if s = ss + 2*n + 6 then + if l = Bk # ires \ r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2*n + 7 then + if r \ [] \ hd r = Oc then 3 + else 0 + else ss+2*n+9 - s)" + +fun abc_inc_stage4 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_inc_stage4 (s, l, r) ss n ires = + (if s \ ss+2*n+1 \ (s - ss) mod 2 = 0 then + if (r\[] \ hd r = Oc) then 0 + else 1 + else if (s \ ss+2*n+1 \ (s - ss) mod 2 = Suc 0) + then length r + else if s = ss + 2*n + 6 then + if l = Bk # ires \ hd r = Bk then 0 + else Suc (length l) + else 0)" + +fun abc_inc_measure :: "(t_conf \ nat \ nat \ block list) \ + (nat \ nat \ nat \ nat)" + where + "abc_inc_measure (c, ss, n, ires) = + (abc_inc_stage1 c ss n, abc_inc_stage2 c ss n, + abc_inc_stage3 c ss n ires, abc_inc_stage4 c ss n ires)" + +definition abc_inc_LE :: "(((nat \ block list \ block list) \ nat \ + nat \ block list) \ ((nat \ block list \ block list) \ nat \ nat \ block list)) set" + where "abc_inc_LE \ (inv_image lex_square abc_inc_measure)" + +lemma wf_lex_triple: "wf lex_triple" +by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def) + +lemma wf_lex_square: "wf lex_square" +by (auto intro:wf_lex_triple simp:lex_triple_def lex_square_def lex_pair_def) + +lemma wf_abc_inc_le[intro]: "wf abc_inc_LE" +by(auto intro:wf_inv_image wf_lex_square simp:abc_inc_LE_def) + +(********************************************************************) +declare inc_inv.simps[simp del] + +lemma halt_lemma2': + "\wf LE; \ n. ((\ P (f n) \ Q (f n)) \ + (Q (f (Suc n)) \ (f (Suc n), (f n)) \ LE)); Q (f 0)\ + \ \ n. P (f n)" +apply(intro exCI, simp) +apply(subgoal_tac "\ n. Q (f n)", simp) +apply(drule_tac f = f in wf_inv_image) +apply(simp add: inv_image_def) +apply(erule wf_induct, simp) +apply(erule_tac x = x in allE) +apply(erule_tac x = n in allE, erule_tac x = n in allE) +apply(erule_tac x = "Suc x" in allE, simp) +apply(rule_tac allI) +apply(induct_tac n, simp) +apply(erule_tac x = na in allE, simp) +done + +lemma halt_lemma2'': + "\P (f n); \ P (f (0::nat))\ \ + \ n. (P (f n) \ (\ i < n. \ P (f i)))" +apply(induct n rule: nat_less_induct, auto) +done + +lemma halt_lemma2''': + "\\n. \ P (f n) \ Q (f n) \ Q (f (Suc n)) \ (f (Suc n), f n) \ LE; + Q (f 0); \i P (f i)\ \ Q (f na)" +apply(induct na, simp, simp) +done + +lemma halt_lemma2: + "\wf LE; + \ n. ((\ P (f n) \ Q (f n)) \ (Q (f (Suc n)) \ (f (Suc n), (f n)) \ LE)); + Q (f 0); \ P (f 0)\ + \ \ n. P (f n) \ Q (f n)" +apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE) +apply(subgoal_tac "\ n. (P (f n) \ (\ i < n. \ P (f i)))") +apply(erule_tac exE)+ +apply(rule_tac x = na in exI, auto) +apply(rule halt_lemma2''', simp, simp, simp) +apply(erule_tac halt_lemma2'', simp) +done + +lemma [simp]: + "\ly = layout_of aprog; abc_fetch as aprog = Some (Inc n)\ + \ start_of ly (Suc as) = start_of ly as + 2*n +9" +apply(case_tac as, auto simp: abc_fetch.simps start_of.simps + layout_of.simps length_of.simps split: if_splits) +done + +lemma inc_inv_init: + "\abc_fetch as aprog = Some (Inc n); + crsp_l ly (as, am) (start_of ly as, l, r) ires; ly = layout_of aprog\ + \ inc_inv ly n (as, am) (start_of ly as, l, r) ires" +apply(auto simp: crsp_l.simps inc_inv.simps + inv_locate_a.simps at_begin_fst_bwtn.simps + at_begin_fst_awtn.simps at_begin_norm.simps ) +apply(auto intro: startof_not0) +done + +lemma inc_inv_stop_pre[simp]: + "\ly = layout_of aprog; inc_inv ly n (as, am) (s, l, r) ires; + s = start_of ly as; abc_fetch as aprog = Some (Inc n)\ + \ (\na. \ (\((s, l, r), ss, n', ires'). s = start_of ly (Suc as)) + (t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) na, s, n, ires) \ + (\((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires') + (t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) na, s, n, ires) \ + (\((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires') + (t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) (Suc na), s, n, ires) \ + ((t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) (Suc na), s, n, ires), + t_steps (s, l, r) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) na, s, n, ires) \ abc_inc_LE)" +apply(rule allI, rule impI, simp add: t_steps_ind, + rule conjI, erule_tac conjE) +apply(rule_tac inc_inv_step, simp, simp, simp) +apply(case_tac "t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na", simp) +proof - + fix na + assume h1: "abc_fetch as aprog = Some (Inc n)" + "\ (\(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) as + 2 * n + 9) + (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n, ires) \ + inc_inv (layout_of aprog) n (as, am) (t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) ires" + from h1 have h2: "start_of (layout_of aprog) as > 0" + apply(rule_tac startof_not0) + done + from h1 and h2 show "((t_step (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0), + start_of (layout_of aprog) as, n, ires), + t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n, ires) + \ abc_inc_LE" + apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), + start_of (layout_of aprog) as - Suc 0) na)", simp) + apply(case_tac "a = 0", + auto split:if_splits simp add:t_step.simps inc_inv.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) + apply(simp_all add:fetch_simps new_tape.simps) + apply(auto simp add: abc_inc_LE_def + lex_square_def lex_triple_def lex_pair_def + inv_after_write.simps inv_after_move.simps inv_after_clear.simps + inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits) + done +qed + +lemma inc_inv_stop_pre1: + "\ + ly = layout_of aprog; + abc_fetch as aprog = Some (Inc n); + s = start_of ly as; + inc_inv ly n (as, am) (s, l, r) ires + \ \ + (\ stp > 0. (\ (s', l', r'). + s' = start_of ly (Suc as) \ + inc_inv ly n (as, am) (s', l', r') ires) + (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp))" +apply(insert halt_lemma2[of abc_inc_LE + "\ ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)" + "(\ stp. (t_steps (s, l, r) + (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp, s, n, ires))" + "\ ((s, l, r), ss, n'). inc_inv ly n (as, am) (s, l, r) ires"]) +apply(insert wf_abc_inc_le) +apply(insert inc_inv_stop_pre[of ly aprog n as am s l r ires], simp) +apply(simp only: t_steps.simps, auto) +apply(rule_tac x = na in exI) +apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Inc n), start_of (layout_of aprog) as - Suc 0) na)", simp) +apply(case_tac na, simp add: t_steps.simps, simp) +done + +lemma inc_inv_stop: + assumes program_and_layout: + -- {* There is an Abacus program @{text "aprog"} and its layout is @{text "ly"}: *} + "ly = layout_of aprog" + and an_instruction: + -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *} + "abc_fetch as aprog = Some (Inc n)" + and the_start_state: + -- {* According to @{text "ly"} and @{text "as"}, + the start state of the TM compiled from this + @{text "Inc n"} instruction should be @{text "s"}: + *} + "s = start_of ly as" + and inv: + -- {* Invariant holds on configuration @{text "(s, l, r)"} *} + "inc_inv ly n (as, am) (s, l, r) ires" + shows -- {* After @{text "stp"} steps of execution, the compiled + TM reaches the start state of next compiled TM and the invariant + still holds. + *} + "(\ stp > 0. (\ (s', l', r'). + s' = start_of ly (Suc as) \ + inc_inv ly n (as, am) (s', l', r') ires) + (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp))" +proof - + from inc_inv_stop_pre1 [OF program_and_layout an_instruction the_start_state inv] + show ?thesis . +qed + +lemma inc_inv_stop_cond: + "\ly = layout_of aprog; + s' = start_of ly (as + 1); + inc_inv ly n (as, lm) (s', (l', r')) ires; + abc_fetch as aprog = Some (Inc n)\ \ + crsp_l ly (Suc as, abc_lm_s lm n (Suc (abc_lm_v lm n))) + (s', l', r') ires" +apply(subgoal_tac "s' = start_of ly as + 2*n + 9", simp) +apply(auto simp: inc_inv.simps inv_stop.simps crsp_l.simps ) +done + +lemma inc_crsp_ex_pre: + "\ly = layout_of aprog; + crsp_l ly (as, am) tc ires; + abc_fetch as aprog = Some (Inc n)\ + \ \stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) + (t_steps tc (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp) ires" +proof(case_tac tc, simp add: abc_step_l.simps) + fix a b c + assume h1: "ly = layout_of aprog" + "crsp_l (layout_of aprog) (as, am) (a, b, c) ires" + "abc_fetch as aprog = Some (Inc n)" + hence h2: "a = start_of ly as" + by(auto simp: crsp_l.simps) + from h1 and h2 have h3: + "inc_inv ly n (as, am) (start_of ly as, b, c) ires" + by(rule_tac inc_inv_init, simp, simp, simp) + from h1 and h2 and h3 have h4: + "(\ stp > 0. (\ (s', l', r'). s' = + start_of ly (Suc as) \ inc_inv ly n (as, am) (s', l', r') ires) + (t_steps (a, b, c) (ci ly (start_of ly as) + (Inc n), start_of ly as - Suc 0) stp))" + apply(rule_tac inc_inv_stop, auto) + done + from h1 and h2 and h3 and h4 show + "\stp > 0. crsp_l (layout_of aprog) + (Suc as, abc_lm_s am n (Suc (abc_lm_v am n))) + (t_steps (a, b, c) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), + start_of (layout_of aprog) as - Suc 0) stp) ires" + apply(erule_tac exE) + apply(rule_tac x = stp in exI) + apply(case_tac "(t_steps (a, b, c) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Inc n), + start_of (layout_of aprog) as - Suc 0) stp)", simp) + apply(rule_tac inc_inv_stop_cond, auto) + done +qed + +text {* + The total correctness of the compilaton of @{text "Inc n"} instruction. +*} + +lemma inc_crsp_ex: + assumes layout: + -- {* For any Abacus program @{text "aprog"}, assuming its layout is @{text "ly"} *} + "ly = layout_of aprog" + and corresponds: + -- {* Abacus configuration @{text "(as, am)"} is in correspondence with + TM configuration @{text "tc"} *} + "crsp_l ly (as, am) tc ires" + and inc: + -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *} + "abc_fetch as aprog = Some (Inc n)" + shows + -- {* + After @{text "stp"} steps of execution, the TM compiled from this @{text "Inc n"} + stops with a configuration which corresponds to the Abacus configuration obtained + from the execution of this @{text "Inc n"} instruction. + *} + "\stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) + (t_steps tc (ci ly (start_of ly as) (Inc n), + start_of ly as - Suc 0) stp) ires" +proof - + from inc_crsp_ex_pre [OF layout corresponds inc] show ?thesis . +qed + +(* +subsection {* The compilation of @{text "Dec n e"} *} +*) + +text {* + The lemmas in this section lead to the correctness of the compilation + of @{text "Dec n e"} instruction using the same techniques as + @{text "Inc n"}. +*} + +type_synonym dec_inv_t = "(nat * nat list) \ t_conf \ block list \ bool" + +fun dec_first_on_right_moving :: "nat \ dec_inv_t" + where + "dec_first_on_right_moving n (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ length lm1 = n \ ml > 0 \ m > 0 \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>)) \ (r = (Oc\<^bsup>mr\<^esup>) \ lm2 = [])))" + +fun dec_on_right_moving :: "dec_inv_t" + where + "dec_on_right_moving (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc (Suc m) \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>)) \ (r = (Oc\<^bsup>mr\<^esup>) \ lm2 = [])))" + +fun dec_after_clear :: "dec_inv_t" + where + "dec_after_clear (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ ml = Suc m \ r \ [] \ r \ [] \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = (Oc\<^bsup>ml \<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + (tl r = Bk # @ (Bk\<^bsup>rn\<^esup>) \ tl r = [] \ lm2 = []))" + +fun dec_after_write :: "dec_inv_t" + where + "dec_after_write (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml + mr = Suc m \ ml = Suc m \ lm2 \ [] \ + (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Bk # (Oc\<^bsup>ml \<^esup>) @ [Bk] @ @ Bk # Bk # ires) \ + tl r = @ (Bk\<^bsup>rn\<^esup>))" + +fun dec_right_move :: "dec_inv_t" + where + "dec_right_move (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 + \ ml = Suc m \ mr = (0::nat) \ + (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Bk # Oc\<^bsup>ml\<^esup>@ [Bk] @ @ Bk # Bk # ires) + \ (r = Bk # @ Bk\<^bsup>rn\<^esup>\ r = [] \ lm2 = []))" + +fun dec_check_right_move :: "dec_inv_t" + where + "dec_check_right_move (as, lm) (s, l, r) ires = + (\ lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \ + ml = Suc m \ mr = (0::nat) \ + (if lm1 = [] then l = Bk # Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Bk # Bk # Oc\<^bsup>ml \<^esup>@ [Bk] @ @ Bk # Bk # ires) \ + r = @ Bk\<^bsup>rn\<^esup>)" + +fun dec_left_move :: "dec_inv_t" + where + "dec_left_move (as, lm) (s, l, r) ires = + (\ lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \ + rn > 0 \ + (if lm1 = [] then l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires + else l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # @ Bk # Bk # ires) \ r = Bk\<^bsup>rn\<^esup>)" + +declare + dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del] + dec_after_write.simps[simp del] dec_left_move.simps[simp del] + dec_check_right_move.simps[simp del] dec_right_move.simps[simp del] + dec_first_on_right_moving.simps[simp del] + +fun inv_locate_n_b :: "inc_inv_t" + where + "inv_locate_n_b (as, lm) (s, l, r) ires= + (\ lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2 \ + length lm1 = s \ m + 1 = ml + mr \ + ml = 1 \ tn = s + 1 - length lm \ + (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires + else l = Oc\<^bsup>ml\<^esup>@Bk#@Bk#Bk#ires) \ + (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ @ (Bk\<^bsup>rn\<^esup>) \ (lm2 = [] \ r = (Oc\<^bsup>mr\<^esup>))) + )" + +fun dec_inv_1 :: "layout \ nat \ nat \ dec_inv_t" + where + "dec_inv_1 ly n e (as, am) (s, l, r) ires = + (let ss = start_of ly as in + let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in + let am'' = abc_lm_s am n (abc_lm_v am n) in + if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires + else if s = ss then False + else if ss \ s \ s < ss + 2*n then + if (s - ss) mod 2 = 0 then + inv_locate_a (as, am) ((s - ss) div 2, l, r) ires + \ inv_locate_a (as, am'') ((s - ss) div 2, l, r) ires + else + inv_locate_b (as, am) ((s - ss) div 2, l, r) ires + \ inv_locate_b (as, am'') ((s - ss) div 2, l, r) ires + else if s = ss + 2 * n then + inv_locate_a (as, am) (n, l, r) ires + \ inv_locate_a (as, am'') (n, l, r) ires + else if s = ss + 2 * n + 1 then + inv_locate_b (as, am) (n, l, r) ires + else if s = ss + 2 * n + 13 then + inv_on_left_moving (as, am'') (s, l, r) ires + else if s = ss + 2 * n + 14 then + inv_check_left_moving (as, am'') (s, l, r) ires + else if s = ss + 2 * n + 15 then + inv_after_left_moving (as, am'') (s, l, r) ires + else False)" + +fun dec_inv_2 :: "layout \ nat \ nat \ dec_inv_t" + where + "dec_inv_2 ly n e (as, am) (s, l, r) ires = + (let ss = start_of ly as in + let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in + let am'' = abc_lm_s am n (abc_lm_v am n) in + if s = 0 then False + else if s = ss then False + else if ss \ s \ s < ss + 2*n then + if (s - ss) mod 2 = 0 then + inv_locate_a (as, am) ((s - ss) div 2, l, r) ires + else inv_locate_b (as, am) ((s - ss) div 2, l, r) ires + else if s = ss + 2 * n then + inv_locate_a (as, am) (n, l, r) ires + else if s = ss + 2 * n + 1 then + inv_locate_n_b (as, am) (n, l, r) ires + else if s = ss + 2 * n + 2 then + dec_first_on_right_moving n (as, am'') (s, l, r) ires + else if s = ss + 2 * n + 3 then + dec_after_clear (as, am') (s, l, r) ires + else if s = ss + 2 * n + 4 then + dec_right_move (as, am') (s, l, r) ires + else if s = ss + 2 * n + 5 then + dec_check_right_move (as, am') (s, l, r) ires + else if s = ss + 2 * n + 6 then + dec_left_move (as, am') (s, l, r) ires + else if s = ss + 2 * n + 7 then + dec_after_write (as, am') (s, l, r) ires + else if s = ss + 2 * n + 8 then + dec_on_right_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 9 then + dec_after_clear (as, am') (s, l, r) ires + else if s = ss + 2 * n + 10 then + inv_on_left_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 11 then + inv_check_left_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 12 then + inv_after_left_moving (as, am') (s, l, r) ires + else if s = ss + 2 * n + 16 then + inv_stop (as, am') (s, l, r) ires + else False)" + +(*begin: dec_fetch lemmas*) + +lemma dec_fetch_locate_a_o: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; start_of ly as > 0; + a - start_of ly as = 2 * q\ + \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (2 * q)) Oc = (R, a + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append Suc_pre) +apply(subgoal_tac "(findnth n ! Suc (4 * q)) = + findnth (Suc q) ! (4 * q + 1)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 1) = + findnth (Suc q) ! (4 * q + 1)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_a_b: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + start_of ly as > 0; + a - start_of ly as = 2 * q\ + \ fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (2 * q)) Bk = (W1, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "(findnth n ! (4 * q)) = + findnth (Suc q) ! (4 * q )") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n !(4 * q + 0) = + findnth (Suc q) ! (4 * q + 0)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_b_o: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + start_of ly as> 0\ + \ fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (a - start_of ly as)) Oc = (R, a)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of ly as) = 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) = + findnth (Suc q) ! (4 * q + 3)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 3) = + findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_b_b: + "\\ a < start_of ly as; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + start_of ly as > 0\ + \ fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (a - start_of ly as)) Bk = (R, a + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append) +apply(subgoal_tac "\ q. (a - start_of ly as) = 2 * q + 1", auto) +apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = + findnth (Suc q) ! (4 * q + 2)") +apply(simp add: findnth.simps nth_append) +apply(subgoal_tac " findnth n ! (4 * q + 2) = + findnth (Suc q) ! (4 * q + 2)", simp) +apply(rule_tac findnth_nth, auto) +done + +lemma dec_fetch_locate_n_a_o: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (2 * n)) Oc + = (R, start_of ly as + 2*n + 1)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_locate_n_a_b: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (2 * n)) Bk + = (W1, start_of ly as + 2*n)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_locate_n_b_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Oc + = (R, start_of ly as + 2*n + 2)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + + +lemma dec_fetch_locate_n_b_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk + = (L, start_of ly as + 2*n + 13)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_first_on_right_move_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n)))) Oc + = (R, start_of ly as + 2*n + 2)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma dec_fetch_first_on_right_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) + (Suc (Suc (Suc (2 * n)))) Bk + = (L, start_of ly as + 2*n + 3)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append tdec_b_def) +done + +lemma [simp]: "fetch x (a + 2 * n) b = fetch x (2 * n + a) b" +thm arg_cong +apply(rule_tac x = "a + 2*n" and y = "2*n + a" in arg_cong, simp) +done + +lemma dec_fetch_first_after_clear_o: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 4) Oc + = (W0, start_of ly as + 2*n + 3)" +apply(auto simp: ci.simps findnth.simps tshift.simps + tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_first_after_clear_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 4) Bk + = (R, start_of ly as + 2*n + 4)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 4= Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_right_move_b: + "start_of ly as > 0 \ fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 5) Bk + = (R, start_of ly as + 2*n + 5)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 5= Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_right_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 6) Bk + = (L, start_of ly as + 2*n + 6)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_right_move_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) (start_of ly as) + (Dec n e)) (2 * n + 6) Oc + = (L, start_of ly as + 2*n + 7)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_left_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 7) Bk + = (L, start_of ly as + 2*n + 10)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 7 = Suc (2*n + 6)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_write_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 8) Bk + = (W1, start_of ly as + 2*n + 7)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_write_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 8) Oc + = (R, start_of ly as + 2*n + 8)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_right_move_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 9) Bk + = (L, start_of ly as + 2*n + 9)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_right_move_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 9) Oc + = (R, start_of ly as + 2*n + 8)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_clear_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 10) Bk + = (R, start_of ly as + 2*n + 4)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_clear_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 10) Oc + = (W0, start_of ly as + 2*n + 9)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 10= Suc (2*n + 9)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move1_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 11) Oc + = (L, start_of ly as + 2*n + 10)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 11= Suc (2*n + 10)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move1_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 11) Bk + = (L, start_of ly as + 2*n + 11)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move1_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 12) Oc + = (L, start_of ly as + 2*n + 10)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 12= Suc (2*n + 11)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move1_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 12) Bk + = (R, start_of ly as + 2*n + 12)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_left_move1_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 13) Bk + = (R, start_of ly as + 2*n + 16)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 13 = Suc (2*n + 12)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move2_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 14) Oc + = (L, start_of ly as + 2*n + 13)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_on_left_move2_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 14) Bk + = (L, start_of ly as + 2*n + 14)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move2_o: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 15) Oc + = (L, start_of ly as + 2*n + 13)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 15 = Suc (2*n + 14)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_check_left_move2_b: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2 * n + 15) Bk + = (R, start_of ly as + 2*n + 15)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 15= Suc (2*n + 14)", simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_after_left_move2_b: + "\ly = layout_of aprog; + abc_fetch as aprog = Some (Dec n e); + start_of ly as > 0\ \ + fetch (ci (layout_of aprog) (start_of ly as) + (Dec n e)) (2 * n + 16) Bk + = (R, start_of ly e)" +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac "2*n + 16 = Suc (2*n + 15)", + simp only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +lemma dec_fetch_next_state: + "start_of ly as > 0 \ + fetch (ci (layout_of aprog) + (start_of ly as) (Dec n e)) (2* n + 17) b + = (Nop, 0)" +apply(case_tac b) +apply(auto simp: ci.simps findnth.simps + tshift.simps tdec_b_def add3_Suc) +apply(subgoal_tac [!] "2*n + 17 = Suc (2*n + 16)", + simp_all only: fetch.simps) +apply(auto simp: nth_of.simps nth_append) +done + +(*End: dec_fetch lemmas*) +lemmas dec_fetch_simps = + dec_fetch_locate_a_o dec_fetch_locate_a_b dec_fetch_locate_b_o + dec_fetch_locate_b_b dec_fetch_locate_n_a_o + dec_fetch_locate_n_a_b dec_fetch_locate_n_b_o + dec_fetch_locate_n_b_b dec_fetch_first_on_right_move_o + dec_fetch_first_on_right_move_b dec_fetch_first_after_clear_b + dec_fetch_first_after_clear_o dec_fetch_right_move_b + dec_fetch_on_right_move_b dec_fetch_on_right_move_o + dec_fetch_after_clear_b dec_fetch_after_clear_o + dec_fetch_check_right_move_b dec_fetch_check_right_move_o + dec_fetch_left_move_b dec_fetch_on_left_move1_b + dec_fetch_on_left_move1_o dec_fetch_check_left_move1_b + dec_fetch_check_left_move1_o dec_fetch_after_left_move1_b + dec_fetch_on_left_move2_b dec_fetch_on_left_move2_o + dec_fetch_check_left_move2_o dec_fetch_check_left_move2_b + dec_fetch_after_left_move2_b dec_fetch_after_write_b + dec_fetch_after_write_o dec_fetch_next_state + +lemma [simp]: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\ + \ \ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, xs) ires" +apply(insert locate_b_2_locate_a[of a ly as n am aaa xs], simp) +done + +lemma [simp]: + "\start_of ly as \ a; + a < start_of ly as + 2 * n; + (a - start_of ly as) mod 2 = Suc 0; + inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\ + \ \ Suc a < start_of ly as + 2 * n \ + inv_locate_a (as, am) (n, Bk # aaa, []) ires" +apply(insert locate_b_2_locate_a_B[of a ly as n am aaa], simp) +done + +(* +lemma [simp]: "a\<^bsup>0\<^esup>=[]" +apply(simp add: exponent_def) +done +*) + +lemma exp_ind: "a\<^bsup>Suc b\<^esup> = a\<^bsup>b\<^esup> @ [a]" +apply(simp only: exponent_def rep_ind) +done + +lemma [simp]: + "inv_locate_b (as, am) (n, l, Oc # r) ires + \ dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n)) + (Suc (Suc (start_of ly as + 2 * n)), Oc # l, r) ires" +apply(simp only: inv_locate_b.simps + dec_first_on_right_moving.simps in_middle.simps + abc_lm_s.simps abc_lm_v.simps) +apply(erule_tac exE)+ +apply(erule conjE)+ +apply(case_tac "n < length am", simp) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc ml" in exI, rule_tac conjI, rule_tac [1-2] impI) +prefer 3 +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(subgoal_tac "Suc n - length am = Suc (n - length am)", + simp only:exponent_def rep_ind, simp) +apply(rule_tac x = "Suc ml" in exI, simp_all) +apply(rule_tac [!] x = "mr - 1" in exI, simp_all) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: + "\inv_locate_b (as, am) (n, l, r) ires; l \ []\ \ + \ inv_on_left_moving_in_middle_B (as, abc_lm_s am n (abc_lm_v am n)) + (s, tl l, hd l # r) ires" +apply(auto simp: inv_locate_b.simps + inv_on_left_moving_in_middle_B.simps in_middle.simps) +apply(case_tac [!] ml, auto split: if_splits) +done + +lemma [simp]: "inv_locate_b (as, am) (n, l, r) ires \ l \ []" +apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits) +done + +lemma [simp]: "\inv_locate_b (as, am) (n, l, Bk # r) ires; n < length am\ + \ inv_on_left_moving_norm (as, am) (s, tl l, hd l # Bk # r) ires" +apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps + in_middle.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "ml - 1" in exI, auto) +apply(rule_tac [!] x = "Suc mr" in exI) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: "\inv_locate_b (as, am) (n, l, Bk # r) ires; \ n < length am\ + \ inv_on_left_moving_norm (as, am @ + replicate (n - length am) 0 @ [0]) (s, tl l, hd l # Bk # r) ires" +apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps + in_middle.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(subgoal_tac "Suc n - length am = Suc (n - length am)", simp only: rep_ind exponent_def, simp_all) +apply(rule_tac x = "Suc mr" in exI, auto) +done + +lemma inv_locate_b_2_on_left_moving[simp]: + "\inv_locate_b (as, am) (n, l, Bk # r) ires\ + \ (l = [] \ inv_on_left_moving (as, + abc_lm_s am n (abc_lm_v am n)) (s, [], Bk # Bk # r) ires) \ + (l \ [] \ inv_on_left_moving (as, + abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires)" +apply(subgoal_tac "l\[]") +apply(subgoal_tac "\ inv_on_left_moving_in_middle_B + (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires") +apply(simp add:inv_on_left_moving.simps + abc_lm_s.simps abc_lm_v.simps split: if_splits, auto) +done + +lemma [simp]: + "inv_locate_b (as, am) (n, l, []) ires \ + inv_locate_b (as, am) (n, l, [Bk]) ires" +apply(auto simp: inv_locate_b.simps in_middle.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = "Suc (length lm1) - length am" in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = ml in exI, rule_tac x = mr in exI) +apply(auto) +done + +lemma nil_2_nil: " = [] \ lm = []" +apply(auto simp: tape_of_nl_abv) +apply(case_tac lm, simp) +apply(case_tac list, auto simp: tape_of_nat_list.simps) +done + +lemma inv_locate_b_2_on_left_moving_b[simp]: + "inv_locate_b (as, am) (n, l, []) ires + \ (l = [] \ inv_on_left_moving (as, + abc_lm_s am n (abc_lm_v am n)) (s, [], [Bk]) ires) \ + (l \ [] \ inv_on_left_moving (as, abc_lm_s am n + (abc_lm_v am n)) (s, tl l, [hd l]) ires)" +apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s]) +apply(simp only: inv_on_left_moving.simps, simp) +apply(subgoal_tac "\ inv_on_left_moving_in_middle_B + (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp) +apply(simp only: inv_on_left_moving_norm.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, rule_tac x = ml in exI, + rule_tac x = mr in exI, simp) +apply(case_tac mr, simp, simp, case_tac nat, auto intro: nil_2_nil) +done + +lemma [simp]: + "\dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\ + \ dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires" +apply(simp only: dec_first_on_right_moving.simps) +apply(erule exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, auto) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: + "dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \ l \ []" +apply(auto simp: dec_first_on_right_moving.simps split: if_splits) +done + +lemma [elim]: + "\\ length lm1 < length am; + am @ replicate (length lm1 - length am) 0 @ [0::nat] = + lm1 @ m # lm2; + 0 < m\ + \ RR" +apply(subgoal_tac "lm2 = []", simp) +apply(drule_tac length_equal, simp) +done + +lemma [simp]: + "\dec_first_on_right_moving n (as, + abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\ +\ dec_after_clear (as, abc_lm_s am n + (abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires" +apply(simp only: dec_first_on_right_moving.simps + dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps) +apply(erule_tac exE)+ +apply(case_tac "n < length am") +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = "m - 1" in exI, auto simp: ) +apply(case_tac [!] mr, auto) +done + +lemma [simp]: + "\dec_first_on_right_moving n (as, + abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\ +\ (l = [] \ dec_after_clear (as, + abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \ + (l \ [] \ dec_after_clear (as, abc_lm_s am n + (abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)" +apply(subgoal_tac "l \ []", + simp only: dec_first_on_right_moving.simps + dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps) +apply(erule_tac exE)+ +apply(case_tac "n < length am", simp) +apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto) +apply(case_tac [1-2] mr, auto) +apply(case_tac [1-2] m, auto simp: dec_first_on_right_moving.simps split: if_splits) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, Oc # r) ires\ + \ dec_after_clear (as, am) (s', l, Bk # r) ires" +apply(auto simp: dec_after_clear.simps) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, Bk # r) ires\ + \ dec_right_move (as, am) (s', Bk # l, r) ires" +apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, []) ires\ + \ dec_right_move (as, am) (s', Bk # l, []) ires" +apply(auto simp: dec_after_clear.simps dec_right_move.simps ) +done + +lemma [simp]: "\rn. a::block\<^bsup>rn\<^esup> = []" +apply(rule_tac x = 0 in exI, simp) +done + +lemma [simp]: "\dec_after_clear (as, am) (s, l, []) ires\ + \ dec_right_move (as, am) (s', Bk # l, [Bk]) ires" +apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits) +done + +lemma [simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False" +apply(auto simp: dec_right_move.simps) +done + +lemma dec_right_move_2_check_right_move[simp]: + "\dec_right_move (as, am) (s, l, Bk # r) ires\ + \ dec_check_right_move (as, am) (s', Bk # l, r) ires" +apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits) +done + +lemma [simp]: + "dec_right_move (as, am) (s, l, []) ires= + dec_right_move (as, am) (s, l, [Bk]) ires" +apply(simp add: dec_right_move.simps) +apply(rule_tac iffI) +apply(erule_tac [!] exE)+ +apply(erule_tac [2] exE) +apply(rule_tac [!] x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac [!] x = m in exI, auto) +apply(auto intro: nil_2_nil) +done + +lemma [simp]: "\dec_right_move (as, am) (s, l, []) ires\ + \ dec_check_right_move (as, am) (s, Bk # l, []) ires" +apply(insert dec_right_move_2_check_right_move[of as am s l "[]" s'], + simp) +done + +lemma [simp]: "dec_check_right_move (as, am) (s, l, r) ires\ l \ []" +apply(auto simp: dec_check_right_move.simps split: if_splits) +done + +lemma [simp]: "\dec_check_right_move (as, am) (s, l, Oc # r) ires\ + \ dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires" +apply(auto simp: dec_check_right_move.simps dec_after_write.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, auto) +done + +lemma [simp]: "\dec_check_right_move (as, am) (s, l, Bk # r) ires\ + \ dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires" +apply(auto simp: dec_check_right_move.simps + dec_left_move.simps inv_after_move.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto) +apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil) +apply(rule_tac x = "Suc rn" in exI) +apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil) +done + +lemma [simp]: "\dec_check_right_move (as, am) (s, l, []) ires\ + \ dec_left_move (as, am) (s', tl l, [hd l]) ires" +apply(auto simp: dec_check_right_move.simps + dec_left_move.simps inv_after_move.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto) +apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil) +done + +lemma [simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False" +apply(auto simp: dec_left_move.simps inv_after_move.simps) +apply(case_tac [!] rn, auto) +done + +lemma [simp]: "dec_left_move (as, am) (s, l, r) ires + \ l \ []" +apply(auto simp: dec_left_move.simps split: if_splits) +done + +lemma tape_of_nl_abv_cons_ex[simp]: + "\lna. Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>ln\<^esup> = @ Bk\<^bsup>lna\<^esup>" +apply(case_tac "lm1=[]", auto simp: tape_of_nl_abv + tape_of_nat_list.simps) +apply(rule_tac x = "ln" in exI, simp) +apply(simp add: tape_of_nat_list_cons exponent_def) +done + +(* +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>ln\<^esup>, Bk # Bk\<^bsup>rn\<^esup>) ires" +apply(simp only: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto) +done +*) +lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires" +apply(simp add: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def) +done + +lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, [Bk]) ires" +apply(simp add: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def) +done + +lemma [simp]: "lm1 \ [] \ + inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', + Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires" +apply(simp only: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto) +done + +lemma [simp]: "lm1 \ [] \ + inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', + Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk # Bk # ires, [Bk]) ires" +apply(simp only: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto) +done + +lemma [simp]: "dec_left_move (as, am) (s, l, Bk # r) ires + \ inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires" +apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits) +done + +(* +lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) + (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # @ Bk\<^bsup>ln\<^esup>, [Bk]) ires" +apply(auto simp: inv_on_left_moving_in_middle_B.simps) +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto) +done +*) + +lemma [simp]: "dec_left_move (as, am) (s, l, []) ires + \ inv_on_left_moving (as, am) (s', tl l, [hd l]) ires" +apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits) +done + +lemma [simp]: "dec_after_write (as, am) (s, l, Oc # r) ires + \ dec_on_right_moving (as, am) (s', Oc # l, r) ires" +apply(auto simp: dec_after_write.simps dec_on_right_moving.simps) +apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, + rule_tac x = "hd lm2" in exI, simp) +apply(rule_tac x = "Suc 0" in exI,rule_tac x = "Suc (hd lm2)" in exI) +apply(case_tac lm2, simp, simp) +apply(case_tac "list = []", + auto simp: tape_of_nl_abv tape_of_nat_list.simps split: if_splits ) +apply(case_tac rn, auto) +apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps) +apply(case_tac rn, auto) +apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto) +apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps) +apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto) +done + +lemma [simp]: "dec_after_write (as, am) (s, l, Bk # r) ires + \ dec_after_write (as, am) (s', l, Oc # r) ires" +apply(auto simp: dec_after_write.simps) +done + +lemma [simp]: "dec_after_write (as, am) (s, aaa, []) ires + \ dec_after_write (as, am) (s', aaa, [Oc]) ires" +apply(auto simp: dec_after_write.simps) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires + \ dec_on_right_moving (as, am) (s', Oc # l, r) ires" +apply(simp only: dec_on_right_moving.simps) +apply(erule_tac exE)+ +apply(erule conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI, + rule_tac x = "mr - 1" in exI, simp) +apply(case_tac mr, auto) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, r) ires\ l \ []" +apply(auto simp: dec_on_right_moving.simps split: if_splits) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires + \ dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires" +apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps) +apply(case_tac [!] mr, auto split: if_splits) +done + +lemma [simp]: "dec_on_right_moving (as, am) (s, l, []) ires + \ dec_after_clear (as, am) (s', tl l, [hd l]) ires" +apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps) +apply(case_tac mr, simp_all split: if_splits) +apply(rule_tac x = lm1 in exI, simp) +done + +lemma start_of_le: "a < b \ start_of ly a \ start_of ly b" +proof(induct b arbitrary: a, simp, case_tac "a = b", simp) + fix b a + show "start_of ly b \ start_of ly (Suc b)" + apply(case_tac "b::nat", + simp add: start_of.simps, simp add: start_of.simps) + done +next + fix b a + assume h1: "\a. a < b \ start_of ly a \ start_of ly b" + "a < Suc b" "a \ b" + hence "a < b" + by(simp) + from h1 and this have h2: "start_of ly a \ start_of ly b" + by(drule_tac h1, simp) + from h2 show "start_of ly a \ start_of ly (Suc b)" + proof - + have "start_of ly b \ start_of ly (Suc b)" + apply(case_tac "b::nat", + simp add: start_of.simps, simp add: start_of.simps) + done + from h2 and this show "start_of ly a \ start_of ly (Suc b)" + by simp + qed +qed + +lemma start_of_dec_length[simp]: + "\abc_fetch a aprog = Some (Dec n e)\ \ + start_of (layout_of aprog) (Suc a) + = start_of (layout_of aprog) a + 2*n + 16" +apply(case_tac a, auto simp: abc_fetch.simps start_of.simps + layout_of.simps length_of.simps + split: if_splits) +done + +lemma start_of_ge: + "\abc_fetch a aprog = Some (Dec n e); a < e\ \ + start_of (layout_of aprog) e > + start_of (layout_of aprog) a + 2*n + 15" +apply(case_tac "e = Suc a", + simp add: start_of.simps abc_fetch.simps layout_of.simps + length_of.simps split: if_splits) +apply(subgoal_tac "Suc a < e", drule_tac a = "Suc a" + and ly = "layout_of aprog" in start_of_le) +apply(subgoal_tac "start_of (layout_of aprog) (Suc a) + = start_of (layout_of aprog) a + 2*n + 16", simp) +apply(rule_tac start_of_dec_length, simp) +apply(arith) +done + +lemma starte_not_equal[simp]: + "\abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\ + \ (start_of ly e \ Suc (Suc (start_of ly as + 2 * n)) \ + start_of ly e \ start_of ly as + 2 * n + 3 \ + start_of ly e \ start_of ly as + 2 * n + 4 \ + start_of ly e \ start_of ly as + 2 * n + 5 \ + start_of ly e \ start_of ly as + 2 * n + 6 \ + start_of ly e \ start_of ly as + 2 * n + 7 \ + start_of ly e \ start_of ly as + 2 * n + 8 \ + start_of ly e \ start_of ly as + 2 * n + 9 \ + start_of ly e \ start_of ly as + 2 * n + 10 \ + start_of ly e \ start_of ly as + 2 * n + 11 \ + start_of ly e \ start_of ly as + 2 * n + 12 \ + start_of ly e \ start_of ly as + 2 * n + 13 \ + start_of ly e \ start_of ly as + 2 * n + 14 \ + start_of ly e \ start_of ly as + 2 * n + 15)" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp) +apply(drule_tac a = as and e = e in start_of_ge, simp, simp) +done + +lemma [simp]: "\abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\ + \ (Suc (Suc (start_of ly as + 2 * n)) \ start_of ly e \ + start_of ly as + 2 * n + 3 \ start_of ly e \ + start_of ly as + 2 * n + 4 \ start_of ly e \ + start_of ly as + 2 * n + 5 \start_of ly e \ + start_of ly as + 2 * n + 6 \ start_of ly e \ + start_of ly as + 2 * n + 7 \ start_of ly e \ + start_of ly as + 2 * n + 8 \ start_of ly e \ + start_of ly as + 2 * n + 9 \ start_of ly e \ + start_of ly as + 2 * n + 10 \ start_of ly e \ + start_of ly as + 2 * n + 11 \ start_of ly e \ + start_of ly as + 2 * n + 12 \ start_of ly e \ + start_of ly as + 2 * n + 13 \ start_of ly e \ + start_of ly as + 2 * n + 14 \ start_of ly e \ + start_of ly as + 2 * n + 15 \ start_of ly e)" +apply(insert starte_not_equal[of as aprog n e ly], + simp del: starte_not_equal) +apply(erule_tac conjE)+ +apply(rule_tac conjI, simp del: starte_not_equal)+ +apply(rule not_sym, simp) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n as)) (Suc 0) Oc = + (R, Suc (start_of (layout_of aprog) as))" + +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append + Suc_pre tdec_b_def) +apply(insert findnth_nth[of 0 n "Suc 0"], simp) +apply(simp add: findnth.simps) +done + +lemma start_of_inj[simp]: + "\abc_fetch as aprog = Some (Dec n e); e \ as; ly = layout_of aprog\ + \ start_of ly as \ start_of ly e" +apply(case_tac "e < as") +apply(case_tac "as", simp, simp) +apply(case_tac "e = nat", simp add: start_of.simps + layout_of.simps length_of.simps) +apply(subgoal_tac "e < length aprog", simp add: length_of.simps + split: abc_inst.splits) +apply(simp add: abc_fetch.simps split: if_splits) +apply(subgoal_tac "e < nat", drule_tac a = e and b = nat + and ly =ly in start_of_le, simp) +apply(subgoal_tac "start_of ly nat < start_of ly (Suc nat)", + simp, simp add: start_of.simps layout_of.simps) +apply(subgoal_tac "nat < length aprog", simp) +apply(case_tac "aprog ! nat", auto simp: length_of.simps) +apply(simp add: abc_fetch.simps split: if_splits) +apply(subgoal_tac "e > as", drule_tac start_of_ge, auto) +done + +lemma [simp]: "\abc_fetch as aprog = Some (Dec n e); e < as\ + \ Suc (start_of (layout_of aprog) e) - + start_of (layout_of aprog) as = 0" +apply(frule_tac ly = "layout_of aprog" in start_of_le, simp) +apply(subgoal_tac "start_of (layout_of aprog) as \ + start_of (layout_of aprog) e", arith) +apply(rule start_of_inj, auto) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); + 0 < start_of (layout_of aprog) as\ + \ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) (Suc (start_of (layout_of aprog) e) - + start_of (layout_of aprog) as) Oc) + = (if e = as then (R, start_of (layout_of aprog) as + 1) + else (Nop, 0))" +apply(auto split: if_splits) +apply(case_tac "e < as", simp add: fetch.simps) +apply(subgoal_tac " e > as") +apply(drule start_of_ge, simp, + auto simp: fetch.simps ci_length nth_of.simps) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) div 2= length_of (Dec n e)") +defer +apply(simp add: ci_length) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) mod 2 = 0", auto simp: length_of.simps) +done + +lemma [simp]: + "start_of (layout_of aprog) as > 0 \ + fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n as)) (Suc 0) Bk + = (W1, start_of (layout_of aprog) as)" +apply(auto simp: ci.simps findnth.simps fetch.simps nth_of.simps + tshift.simps nth_append Suc_pre tdec_b_def) +apply(insert findnth_nth[of 0 n "0"], simp) +apply(simp add: findnth.simps) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); + 0 < start_of (layout_of aprog) as\ +\ (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) (Suc (start_of (layout_of aprog) e) - + start_of (layout_of aprog) as) Bk) + = (if e = as then (W1, start_of (layout_of aprog) as) + else (Nop, 0))" +apply(auto split: if_splits) +apply(case_tac "e < as", simp add: fetch.simps) +apply(subgoal_tac " e > as") +apply(drule start_of_ge, simp, auto simp: fetch.simps + ci_length nth_of.simps) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) div 2= length_of (Dec n e)") +defer +apply(simp add: ci_length) +apply(subgoal_tac + "length (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) mod 2 = 0", auto simp: length_of.simps) +apply(simp add: ci.simps tshift.simps tdec_b_def) +done + +lemma [simp]: + "inv_stop (as, abc_lm_s am n (abc_lm_v am n)) (s, l, r) ires \ l \ []" +apply(auto simp: inv_stop.simps) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); e \ as; ly = layout_of aprog\ + \ (\ (start_of ly as \ start_of ly e \ + start_of ly e < start_of ly as + 2 * n)) + \ start_of ly e \ start_of ly as + 2*n \ + start_of ly e \ Suc (start_of ly as + 2*n) " +apply(case_tac "e < as") +apply(drule_tac ly = ly in start_of_le, simp) +apply(case_tac n, simp, drule start_of_inj, simp, simp, simp, simp) +apply(drule_tac start_of_ge, simp, simp) +done + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); start_of ly as \ s; + s < start_of ly as + 2 * n; ly = layout_of aprog\ + \ Suc s \ start_of ly e " +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [simp]: "\abc_fetch as aprog = Some (Dec n e); + ly = layout_of aprog\ + \ Suc (start_of ly as + 2 * n) \ start_of ly e" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma dec_false_1[simp]: + "\abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\ + \ False" +apply(auto simp: inv_locate_b.simps in_middle.simps exponent_def) +apply(case_tac "length lm1 \ length am", auto) +apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp) +apply(case_tac mr, auto simp: ) +apply(subgoal_tac "Suc (length lm1) - length am = + Suc (length lm1 - length am)", + simp add: rep_ind del: replicate.simps, simp) +apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0" + and ys = "lm1 @ m # lm2" in length_equal, simp) +apply(case_tac mr, auto simp: abc_lm_v.simps) +apply(case_tac "mr = 0", simp_all add: exponent_def split: if_splits) +apply(subgoal_tac "Suc (length lm1) - length am = + Suc (length lm1 - length am)", + simp add: rep_ind del: replicate.simps, simp) +done + +lemma [simp]: + "\inv_locate_b (as, am) (n, aaa, Bk # xs) ires; + abc_lm_v am n = 0\ + \ inv_on_left_moving (as, abc_lm_s am n 0) + (s, tl aaa, hd aaa # Bk # xs) ires" +apply(insert inv_locate_b_2_on_left_moving[of as am n aaa xs ires s], simp) +done + +lemma [simp]: + "\abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\ + \ inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires" +apply(insert inv_locate_b_2_on_left_moving_b[of as am n aaa ires s], simp) +done + +lemma [simp]: "\am ! n = (0::nat); n < length am\ \ am[n := 0] = am" +apply(simp add: list_update_same_conv) +done + +lemma [simp]: "\abc_lm_v am n = 0; + inv_locate_b (as, abc_lm_s am n 0) (n, Oc # aaa, xs) ires\ + \ inv_locate_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp only: inv_locate_b.simps in_middle.simps abc_lm_s.simps + abc_lm_v.simps) +apply(erule_tac exE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, simp) +apply(case_tac "n < length am", simp_all) +apply(erule_tac conjE)+ +apply(rule_tac x = tn in exI, rule_tac x = m in exI, simp) +apply(rule_tac x = ml in exI, rule_tac x = mr in exI, simp) +defer +apply(rule_tac x = "Suc n - length am" in exI, rule_tac x = m in exI) +apply(subgoal_tac "Suc n - length am = Suc (n - length am)") +apply(simp add: exponent_def rep_ind del: replicate.simps, auto) +done + +lemma [intro]: "\abc_lm_v (a # list) 0 = 0\ \ a = 0" +apply(simp add: abc_lm_v.simps split: if_splits) +done + +lemma [simp]: + "inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires + \ inv_locate_a (as, abc_lm_s am n 0) (0, aaa, Oc # xs) ires" +apply(simp add: inv_locate_a.simps) +apply(rule disjI1) +apply(auto simp: inv_stop.simps at_begin_norm.simps) +done + +lemma [simp]: + "\abc_lm_v am 0 = 0; + inv_stop (as, abc_lm_s am 0 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires\ \ + inv_locate_b (as, am) (0, Oc # aaa, xs) ires" +apply(auto simp: inv_stop.simps inv_locate_b.simps + in_middle.simps abc_lm_s.simps) +apply(case_tac "am = []", simp) +apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, + rule_tac x = 0 in exI, simp) +apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, + simp add: tape_of_nl_abv tape_of_nat_list.simps, auto) +apply(case_tac rn, auto) +apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI, + rule_tac x = "hd am" in exI, simp) +apply(rule_tac x = "Suc 0" in exI, rule_tac x = "hd am" in exI, simp) +apply(case_tac am, simp, simp) +apply(subgoal_tac "a = 0", case_tac list, + auto simp: tape_of_nat_list.simps tape_of_nl_abv) +apply(case_tac rn, auto) +done + +lemma [simp]: + "\inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires\ + \ inv_locate_b (as, am) (0, Oc # aaa, xs) ires \ + inv_locate_b (as, abc_lm_s am n 0) (0, Oc # aaa, xs) ires" +apply(simp) +done + +lemma [simp]: +"\abc_lm_v am n = 0; + inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Oc # xs) ires\ + \ \ Suc 0 < 2 * n \ e = as \ + inv_locate_b (as, am) (n, Oc # aaa, xs) ires" +apply(case_tac n, simp, simp) +done + +lemma dec_false2: + "inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, Bk # xs) ires = False" +apply(auto simp: inv_stop.simps abc_lm_s.simps) +apply(case_tac "am", simp, case_tac n, simp add: tape_of_nl_abv) +apply(case_tac list, simp add: tape_of_nat_list.simps ) +apply(simp add: tape_of_nat_list.simps , simp) +apply(case_tac "list[nat := 0]", + simp add: tape_of_nat_list.simps tape_of_nl_abv) +apply(simp add: tape_of_nat_list.simps ) +apply(case_tac "am @ replicate (n - length am) 0 @ [0]", simp) +apply(case_tac list, auto simp: tape_of_nl_abv + tape_of_nat_list.simps ) +done + +lemma dec_false3: + "inv_stop (as, abc_lm_s am n 0) + (start_of (layout_of aprog) e, aaa, []) ires = False" +apply(auto simp: inv_stop.simps abc_lm_s.simps) +apply(case_tac "am", case_tac n, auto) +apply(case_tac n, auto simp: tape_of_nl_abv) +apply(case_tac "list::nat list", + simp add: tape_of_nat_list.simps tape_of_nat_list.simps) +apply(simp add: tape_of_nat_list.simps) +apply(case_tac "list[nat := 0]", + simp add: tape_of_nat_list.simps tape_of_nat_list.simps) +apply(simp add: tape_of_nat_list.simps) +apply(case_tac "(am @ replicate (n - length am) 0 @ [0])", simp) +apply(case_tac list, auto simp: tape_of_nat_list.simps) +done + +lemma [simp]: + "fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e)) 0 b = (Nop, 0)" +by(simp add: fetch.simps) + +declare dec_inv_1.simps[simp del] + +declare inv_locate_n_b.simps [simp del] + +lemma [simp]: +"\0 < abc_lm_v am n; 0 < n; + at_begin_norm (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp only: at_begin_norm.simps inv_locate_n_b.simps) +apply(erule_tac exE)+ +apply(rule_tac x = lm1 in exI, simp) +apply(case_tac "length lm2", simp) +apply(case_tac rn, simp, simp) +apply(rule_tac x = "tl lm2" in exI, rule_tac x = "hd lm2" in exI, simp) +apply(rule conjI) +apply(case_tac "lm2", simp, simp) +apply(case_tac "lm2", auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac [!] "list", auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, auto) +done +lemma [simp]: "(\rn. Oc # xs = Bk\<^bsup>rn\<^esup>) = False" +apply(auto) +apply(case_tac rn, auto simp: ) +done + +lemma [simp]: + "\0 < abc_lm_v am n; 0 < n; + at_begin_fst_bwtn (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp add: at_begin_fst_bwtn.simps inv_locate_n_b.simps ) +done + +lemma Suc_minus:"length am + tn = n + \ Suc tn = Suc n - length am " +apply(arith) +done + +lemma [simp]: + "\0 < abc_lm_v am n; 0 < n; + at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" +apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps ) +apply(erule exE)+ +apply(erule conjE)+ +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI) +apply(simp add: exponent_def rep_ind del: replicate.simps) +apply(rule conjI)+ +apply(auto) +apply(case_tac [!] rn, auto) +done + +lemma [simp]: + "\0 < abc_lm_v am n; 0 < n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\ + \ inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires" +apply(auto simp: inv_locate_a.simps) +done + +lemma [simp]: + "\inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\ + \ dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n)) + (s, Oc # aaa, xs) ires" +apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps + abc_lm_s.simps abc_lm_v.simps) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc (Suc 0)" in exI, + rule_tac x = "m - 1" in exI, simp) +apply(case_tac m, auto simp: exponent_def) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, + simp add: Suc_diff_le rep_ind del: replicate.simps) +apply(rule_tac x = "Suc (Suc 0)" in exI, + rule_tac x = "m - 1" in exI, simp) +apply(case_tac m, auto simp: exponent_def) +apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, + rule_tac x = m in exI, simp) +apply(rule_tac x = "Suc (Suc 0)" in exI, + rule_tac x = "m - 1" in exI, simp) +apply(case_tac m, auto) +apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, + rule_tac x = m in exI, + simp add: Suc_diff_le rep_ind del: replicate.simps, simp) +done + +lemma dec_false_2: + "\0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\ + \ False" +apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits) +apply(case_tac [!] m, auto) +done + +lemma dec_false_2_b: + "\0 < abc_lm_v am n; inv_locate_n_b (as, am) + (n, aaa, []) ires\ \ False" +apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits) +apply(case_tac [!] m, auto simp: ) +done + + +(*begin: dec halt1 lemmas*) +thm abc_inc_stage1.simps +fun abc_dec_1_stage1:: "t_conf \ nat \ nat \ nat" + where + "abc_dec_1_stage1 (s, l, r) ss n = + (if s > ss \ s \ ss + 2*n + 1 then 4 + else if s = ss + 2 * n + 13 \ s = ss + 2*n + 14 then 3 + else if s = ss + 2*n + 15 then 2 + else 0)" + +fun abc_dec_1_stage2:: "t_conf \ nat \ nat \ nat" + where + "abc_dec_1_stage2 (s, l, r) ss n = + (if s \ ss + 2 * n + 1 then (ss + 2 * n + 16 - s) + else if s = ss + 2*n + 13 then length l + else if s = ss + 2*n + 14 then length l + else 0)" + +fun abc_dec_1_stage3 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_dec_1_stage3 (s, l, r) ss n ires = + (if s \ ss + 2*n + 1 then + if (s - ss) mod 2 = 0 then + if r \ [] \ hd r = Oc then 0 else 1 + else length r + else if s = ss + 2 * n + 13 then + if l = Bk # ires \ r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2 * n + 14 then + if r \ [] \ hd r = Oc then 3 else 0 + else 0)" + +fun abc_dec_1_measure :: "(t_conf \ nat \ nat \ block list) \ (nat \ nat \ nat)" + where + "abc_dec_1_measure (c, ss, n, ires) = (abc_dec_1_stage1 c ss n, + abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n ires)" + +definition abc_dec_1_LE :: + "(((nat \ block list \ block list) \ nat \ + nat \ block list) \ ((nat \ block list \ block list) \ nat \ nat \ block list)) set" + where "abc_dec_1_LE \ (inv_image lex_triple abc_dec_1_measure)" + +lemma wf_dec_le: "wf abc_dec_1_LE" +by(auto intro:wf_inv_image wf_lex_triple simp:abc_dec_1_LE_def) + +declare dec_inv_1.simps[simp del] dec_inv_2.simps[simp del] + +lemma [elim]: + "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) as < start_of (layout_of aprog) e; + start_of (layout_of aprog) e \ + Suc (start_of (layout_of aprog) as + 2 * n)\ \ False" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "layout_of aprog" in + start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [elim]: "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) e + = start_of (layout_of aprog) as + 2 * n + 13\ + \ False" +apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], + simp) +done + +lemma [elim]: "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) e = + start_of (layout_of aprog) as + 2 * n + 14\ + \ False" +apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], + simp) +done + +lemma [elim]: + "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) as < start_of (layout_of aprog) e; + start_of (layout_of aprog) e \ + Suc (start_of (layout_of aprog) as + 2 * n)\ + \ False" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "layout_of aprog" in + start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [elim]: + "\abc_fetch as aprog = Some (Dec n e); + start_of (layout_of aprog) e = + start_of (layout_of aprog) as + 2 * n + 13\ + \ False" +apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], + simp) +done + +lemma [simp]: + "abc_fetch as aprog = Some (Dec n e) \ + Suc (start_of (layout_of aprog) as) \ start_of (layout_of aprog) e" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "(layout_of aprog)" in + start_of_le, simp) +apply(drule_tac a = as and e = e in start_of_ge, simp, simp) +done + +lemma [simp]: "inv_on_left_moving (as, am) (s, [], r) ires + = False" +apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps + inv_on_left_moving_in_middle_B.simps) +done + +lemma [simp]: + "inv_check_left_moving (as, abc_lm_s am n 0) + (start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires + = False" +apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps) +done + +lemma dec_inv_stop1_pre: + "\abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0; + start_of (layout_of aprog) as > 0\ + \ \na. \ (\(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) e) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n, ires) \ + dec_inv_1 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na) ires + \ dec_inv_1 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) + (Suc na)) ires \ + ((t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na), + start_of (layout_of aprog) as, n, ires), + t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n, ires) + \ abc_dec_1_LE" +apply(rule allI, rule impI, simp add: t_steps_ind) +apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) +(ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), +start_of (layout_of aprog) as - Suc 0) na)", simp) +apply(auto split:if_splits simp add:t_step.simps dec_inv_1.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) +apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_1.simps) +apply(auto simp add: abc_dec_1_LE_def lex_square_def + lex_triple_def lex_pair_def + split: if_splits) +apply(rule dec_false_1, simp, simp) +done + +lemma dec_inv_stop1: + "\ly = layout_of aprog; + dec_inv_1 ly n e (as, am) (start_of ly as + 1, l, r) ires; + abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0\ \ + (\ stp. (\ (s', l', r'). s' = start_of ly e \ + dec_inv_1 ly n e (as, am) (s', l' , r') ires) + (t_steps (start_of ly as + 1, l, r) + (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp))" +apply(insert halt_lemma2[of abc_dec_1_LE + "\ ((s, l, r), ss, n', ires'). s = start_of ly e" + "(\ stp. (t_steps (start_of ly as + 1, l, r) + (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) + stp, start_of ly as, n, ires))" + "\ ((s, l, r), ss, n, ires'). dec_inv_1 ly n e (as, am) (s, l, r) ires'"], + simp) +apply(insert wf_dec_le, simp) +apply(insert dec_inv_stop1_pre[of as aprog n e am l r], simp) +apply(subgoal_tac "start_of (layout_of aprog) as > 0", + simp add: t_steps.simps) +apply(erule_tac exE, rule_tac x = na in exI) +apply(case_tac + "(t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) na)", + case_tac b, auto) +apply(rule startof_not0) +done + +(*begin: dec halt2 lemmas*) + +lemma [simp]: + "\abc_fetch as aprog = Some (Dec n e); + ly = layout_of aprog\ \ + start_of ly (Suc as) = start_of ly as + 2*n + 16" +by simp + +fun abc_dec_2_stage1 :: "t_conf \ nat \ nat \ nat" + where + "abc_dec_2_stage1 (s, l, r) ss n = + (if s \ ss + 2*n + 1 then 7 + else if s = ss + 2*n + 2 then 6 + else if s = ss + 2*n + 3 then 5 + else if s \ ss + 2*n + 4 \ s \ ss + 2*n + 9 then 4 + else if s = ss + 2*n + 6 then 3 + else if s = ss + 2*n + 10 \ s = ss + 2*n + 11 then 2 + else if s = ss + 2*n + 12 then 1 + else 0)" + +thm new_tape.simps + +fun abc_dec_2_stage2 :: "t_conf \ nat \ nat \ nat" + where + "abc_dec_2_stage2 (s, l, r) ss n = + (if s \ ss + 2 * n + 1 then (ss + 2 * n + 16 - s) + else if s = ss + 2*n + 10 then length l + else if s = ss + 2*n + 11 then length l + else if s = ss + 2*n + 4 then length r - 1 + else if s = ss + 2*n + 5 then length r + else if s = ss + 2*n + 7 then length r - 1 + else if s = ss + 2*n + 8 then + length r + length (takeWhile (\ a. a = Oc) l) - 1 + else if s = ss + 2*n + 9 then + length r + length (takeWhile (\ a. a = Oc) l) - 1 + else 0)" + +fun abc_dec_2_stage3 :: "t_conf \ nat \ nat \ block list \ nat" + where + "abc_dec_2_stage3 (s, l, r) ss n ires = + (if s \ ss + 2*n + 1 then + if (s - ss) mod 2 = 0 then if r \ [] \ + hd r = Oc then 0 else 1 + else length r + else if s = ss + 2 * n + 10 then + if l = Bk # ires \ r \ [] \ hd r = Oc then 2 + else 1 + else if s = ss + 2 * n + 11 then + if r \ [] \ hd r = Oc then 3 + else 0 + else (ss + 2 * n + 16 - s))" + +fun abc_dec_2_stage4 :: "t_conf \ nat \ nat \ nat" + where + "abc_dec_2_stage4 (s, l, r) ss n = + (if s = ss + 2*n + 2 then length r + else if s = ss + 2*n + 8 then length r + else if s = ss + 2*n + 3 then + if r \ [] \ hd r = Oc then 1 + else 0 + else if s = ss + 2*n + 7 then + if r \ [] \ hd r = Oc then 0 + else 1 + else if s = ss + 2*n + 9 then + if r \ [] \ hd r = Oc then 1 + else 0 + else 0)" + +fun abc_dec_2_measure :: "(t_conf \ nat \ nat \ block list) \ + (nat \ nat \ nat \ nat)" + where + "abc_dec_2_measure (c, ss, n, ires) = + (abc_dec_2_stage1 c ss n, abc_dec_2_stage2 c ss n, + abc_dec_2_stage3 c ss n ires, abc_dec_2_stage4 c ss n)" + +definition abc_dec_2_LE :: + "(((nat \ block list \ block list) \ nat \ nat \ block list) \ + ((nat \ block list \ block list) \ nat \ nat \ block list)) set" + where "abc_dec_2_LE \ (inv_image lex_square abc_dec_2_measure)" + +lemma wf_dec_2_le: "wf abc_dec_2_LE" +by(auto intro:wf_inv_image wf_lex_triple wf_lex_square + simp:abc_dec_2_LE_def) + +lemma [simp]: "dec_after_write (as, am) (s, aa, r) ires + \ takeWhile (\a. a = Oc) aa = []" +apply(simp only : dec_after_write.simps) +apply(erule exE)+ +apply(erule_tac conjE)+ +apply(case_tac aa, simp) +apply(case_tac a, simp only: takeWhile.simps , simp, simp split: if_splits) +done + +lemma [simp]: + "\dec_on_right_moving (as, lm) (s, aa, []) ires; + length (takeWhile (\a. a = Oc) (tl aa)) + \ length (takeWhile (\a. a = Oc) aa) - Suc 0\ + \ length (takeWhile (\a. a = Oc) (tl aa)) < + length (takeWhile (\a. a = Oc) aa) - Suc 0" +apply(simp only: dec_on_right_moving.simps) +apply(erule_tac exE)+ +apply(erule_tac conjE)+ +apply(case_tac mr, auto split: if_splits) +done + +lemma [simp]: + "dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires + \ length xs - Suc 0 < length xs + + length (takeWhile (\a. a = Oc) aa)" +apply(simp only: dec_after_clear.simps) +apply(erule_tac exE)+ +apply(erule conjE)+ +apply(simp split: if_splits ) +done + +lemma [simp]: + "\dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\ + \ Suc 0 < length (takeWhile (\a. a = Oc) aa)" +apply(simp add: dec_after_clear.simps split: if_splits) +done + +lemma [simp]: + "\dec_on_right_moving (as, am) (s, aa, Bk # xs) ires; + Suc (length (takeWhile (\a. a = Oc) (tl aa))) + \ length (takeWhile (\a. a = Oc) aa)\ + \ Suc (length (takeWhile (\a. a = Oc) (tl aa))) + < length (takeWhile (\a. a = Oc) aa)" +apply(simp only: dec_on_right_moving.simps) +apply(erule exE)+ +apply(erule conjE)+ +apply(case_tac ml, auto split: if_splits ) +done + +(* +lemma abc_dec_2_wf: + "\ly = layout_of aprog; dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r); abc_fetch as aprog = Dec n e; abc_lm_v am n > 0\ + \ \na. \ (\(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n) \ + ((t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na), + start_of (layout_of aprog) as, n), + t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n) + \ abc_dec_2_LE" +proof(rule allI, rule impI, simp add: t_steps_ind) + fix na + assume h1 :"ly = layout_of aprog" "dec_inv_2 (layout_of aprog) n e (as, am) (Suc (start_of (layout_of aprog) as), l, r)" + "abc_fetch as aprog = Dec n e" "abc_lm_v am n > 0" + "\ (\(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n)" + thus "((t_step (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0), + start_of (layout_of aprog) as, n), + t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n) + \ abc_dec_2_LE" + proof(insert dec_inv_2_steps[of "layout_of aprog" n e as am "(start_of (layout_of aprog) as + 1, l, r)" aprog na], + case_tac "(t_steps (start_of (layout_of aprog) as + 1, l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)", case_tac b, simp) + fix a b aa ba + assume "dec_inv_2 (layout_of aprog) n e (as, am) (a, aa, ba)" " a \ start_of (layout_of aprog) as + 2*n + 16" "abc_lm_v am n > 0" "abc_fetch as aprog = Dec n e " + thus "((t_step (a, aa, ba) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0), start_of (layout_of aprog) as, n), (a, aa, ba), + start_of (layout_of aprog) as, n) + \ abc_dec_2_LE" + apply(case_tac "a = 0", auto split:if_splits simp add:t_step.simps dec_inv_2.simps, + tactic {* ALLGOALS (resolve_tac (thms "fetch_intro")) *}) + apply(simp_all add:dec_fetch_simps new_tape.simps) + apply(auto simp add: abc_dec_2_LE_def lex_square_def lex_triple_def lex_pair_def + split: if_splits) + + done + qed +qed +*) + +lemma [simp]: "inv_check_left_moving (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (start_of (layout_of aprog) as + 2 * n + 11, [], Oc # xs) ires = False" +apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps) +done + +lemma dec_inv_stop2_pre: + "\abc_fetch as aprog = Some (Dec n e); abc_lm_v am n > 0\ \ + \na. \ (\(s, l, r) (ss, n', ires'). + s = start_of (layout_of aprog) as + 2 * n + 16) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na) + (start_of (layout_of aprog) as, n, ires) \ + dec_inv_2 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na) ires + \ + dec_inv_2 (layout_of aprog) n e (as, am) + (t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) (Suc na)) ires \ + ((t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) (Suc na), + start_of (layout_of aprog) as, n, ires), + t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na, + start_of (layout_of aprog) as, n, ires) + \ abc_dec_2_LE" +apply(subgoal_tac "start_of (layout_of aprog) as > 0") +apply(rule allI, rule impI, simp add: t_steps_ind) +apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na)", simp) +apply(auto split:if_splits simp add:t_step.simps dec_inv_2.simps, + tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) +apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_2.simps) +apply(auto simp add: abc_dec_2_LE_def lex_square_def lex_triple_def + lex_pair_def split: if_splits) +apply(auto intro: dec_false_2_b dec_false_2) +apply(rule startof_not0) +done + +lemma dec_stop2: + "\ly = layout_of aprog; + dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r) ires; + abc_fetch as aprog = Some (Dec n e); + abc_lm_v am n > 0\ \ + (\ stp. (\ (s', l', r'). s' = start_of ly (Suc as) \ + dec_inv_2 ly n e (as, am) (s', l', r') ires) + (t_steps (start_of ly as+1, l, r) (ci ly (start_of ly as) + (Dec n e), start_of ly as - Suc 0) stp))" +apply(insert halt_lemma2[of abc_dec_2_LE + "\ ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)" + "(\ stp. (t_steps (start_of ly as + 1, l, r) + (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp, + start_of ly as, n, ires))" + "(\ ((s, l, r), ss, n, ires'). dec_inv_2 ly n e (as, am) (s, l, r) ires')"]) +apply(insert wf_dec_2_le, simp) +apply(insert dec_inv_stop2_pre[of as aprog n e am l r], + simp add: t_steps.simps) +apply(erule_tac exE) +apply(rule_tac x = na in exI) +apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) +(ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) na)", + case_tac b, auto) +done + +lemma dec_inv_stop_cond1: + "\ly = layout_of aprog; + dec_inv_1 ly n e (as, lm) (s, (l, r)) ires; s = start_of ly e; + abc_fetch as aprog = Some (Dec n e); abc_lm_v lm n = 0\ + \ crsp_l ly (e, abc_lm_s lm n 0) (s, l, r) ires" +apply(simp add: dec_inv_1.simps split: if_splits) +apply(auto simp: crsp_l.simps inv_stop.simps ) +done + +lemma dec_inv_stop_cond2: + "\ly = layout_of aprog; s = start_of ly (Suc as); + dec_inv_2 ly n e (as, lm) (s, (l, r)) ires; + abc_fetch as aprog = Some (Dec n e); + abc_lm_v lm n > 0\ + \ crsp_l ly (Suc as, + abc_lm_s lm n (abc_lm_v lm n - Suc 0)) (s, l, r) ires" +apply(simp add: dec_inv_2.simps split: if_splits) +apply(auto simp: crsp_l.simps inv_stop.simps ) +done + +lemma [simp]: "(case Bk\<^bsup>rn\<^esup> of [] \ Bk | + Bk # xs \ Bk | Oc # xs \ Oc) = Bk" +apply(case_tac rn, auto) +done + +lemma [simp]: "t_steps tc (p,off) (m + n) = + t_steps (t_steps tc (p, off) m) (p, off) n" +apply(induct m arbitrary: n) +apply(simp add: t_steps.simps) +proof - + fix m n + assume h1: "\n. t_steps tc (p, off) (m + n) = + t_steps (t_steps tc (p, off) m) (p, off) n" + hence h2: "t_steps tc (p, off) (Suc m + n) = + t_steps tc (p, off) (m + Suc n)" + by simp + from h1 and this show + "t_steps tc (p, off) (Suc m + n) = + t_steps (t_steps tc (p, off) (Suc m)) (p, off) n" + proof(simp only: h2, simp add: t_steps.simps) + have h3: "(t_step (t_steps tc (p, off) m) (p, off)) = + (t_steps (t_step tc (p, off)) (p, off) m)" + apply(simp add: t_steps.simps[THEN sym] t_steps_ind[THEN sym]) + done + from h3 show + "t_steps (t_step (t_steps tc (p, off) m) (p, off)) (p, off) n = t_steps (t_steps (t_step tc (p, off)) (p, off) m) (p, off) n" + by simp + qed +qed + +lemma [simp]: " abc_fetch as aprog = Some (Dec n e) \ + Suc (start_of (layout_of aprog) as) \ + start_of (layout_of aprog) e" +apply(case_tac "e = as", simp) +apply(case_tac "e < as") +apply(drule_tac a = e and b = as and ly = "layout_of aprog" + in start_of_le, simp) +apply(drule_tac start_of_ge, auto) +done + +lemma [simp]: "inv_locate_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires" +apply(auto simp: inv_locate_b.simps in_middle.simps) +apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, + rule_tac x = 0 in exI, simp) +apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, auto) +apply(case_tac rn, simp, case_tac nat, auto) +done + +lemma [simp]: + "inv_locate_n_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires" +apply(auto simp: inv_locate_n_b.simps in_middle.simps) +apply(case_tac rn, simp, case_tac nat, auto) +done + +lemma [simp]: +"abc_fetch as aprog = Some (Dec n e) \ + dec_inv_1 (layout_of aprog) n e (as, []) + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires +\ + dec_inv_2 (layout_of aprog) n e (as, []) + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires" +apply(simp add: dec_inv_1.simps dec_inv_2.simps) +apply(case_tac n, auto) +done + +lemma [simp]: + "\am \ []; = Oc # r'; + abc_fetch as aprog = Some (Dec n e)\ + \ inv_locate_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires" +apply(auto simp: inv_locate_b.simps in_middle.simps) +apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI, + rule_tac x = "hd am" in exI, simp) +apply(rule_tac x = "Suc 0" in exI) +apply(rule_tac x = "hd am" in exI, simp) +apply(case_tac am, simp, case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, auto) +done + +lemma [simp]: + "\ = Oc # r'; abc_fetch as aprog = Some (Dec n e)\ \ + inv_locate_n_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires" +apply(auto simp: inv_locate_n_b.simps) +apply(rule_tac x = "tl am" in exI, rule_tac x = "hd am" in exI, auto) +apply(case_tac [!] am, auto simp: tape_of_nl_abv tape_of_nat_list.simps ) +apply(case_tac [!]list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +apply(case_tac rn, simp, simp) +apply(erule_tac x = nat in allE, simp) +done + +lemma [simp]: + "\am \ []; + = Oc # r'; + abc_fetch as aprog = Some (Dec n e)\ \ + dec_inv_1 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), + Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires \ + dec_inv_2 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), + Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires" +apply(simp add: dec_inv_1.simps dec_inv_2.simps) +apply(case_tac n, auto) +done + +lemma [simp]: "am \ [] \ \r'. = Oc # r'" +apply(case_tac am, simp, case_tac list) +apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps ) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + (fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e)) (Suc 0) Bk) + = (W1, start_of (layout_of aprog) as)" +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append Suc_pre tdec_b_def) +thm findnth_nth +apply(insert findnth_nth[of 0 n 0], simp) +apply(simp add: findnth.simps) +done + +lemma [simp]: + "start_of (layout_of aprog) as > 0 + \ (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Bk\<^bsup>rn\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0)) + = (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn- Suc 0\<^esup>)" +apply(simp add: t_step.simps) +apply(case_tac "start_of (layout_of aprog) as", + auto simp: new_tape.simps) +apply(case_tac rn, auto) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e)) (Suc 0) Oc) + = (R, Suc (start_of (layout_of aprog) as))" + +apply(auto simp: ci.simps findnth.simps fetch.simps + nth_of.simps tshift.simps nth_append + Suc_pre tdec_b_def) +apply(insert findnth_nth[of 0 n "Suc 0"], simp) +apply(simp add: findnth.simps) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn - Suc 0\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0)) = + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn-Suc 0\<^esup>)" +apply(simp add: t_step.simps) +apply(case_tac "start_of (layout_of aprog) as", + auto simp: new_tape.simps) +done + +lemma [simp]: "start_of (layout_of aprog) as > 0 \ + t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # r' @ Bk\<^bsup>rn\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) = + (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>)" +apply(simp add: t_step.simps) +apply(case_tac "start_of (layout_of aprog) as", + auto simp: new_tape.simps) +done + +lemma crsp_next_state: + "\crsp_l (layout_of aprog) (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e)\ + \ \ stp' > 0. (\ (s, l, r). + (s = Suc (start_of (layout_of aprog) as) + \ (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r) ires) + \ (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')" +apply(subgoal_tac "start_of (layout_of aprog) as > 0") +apply(case_tac tc, case_tac b, auto simp: crsp_l.simps) +apply(case_tac "am = []", simp) +apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: t_steps.simps) +proof- + fix rn + assume h1: "am \ []" "abc_fetch as aprog = Some (Dec n e)" + "start_of (layout_of aprog) as > 0" + hence h2: "\ r'. = Oc # r'" + by simp + from h1 and h2 show + "\stp'>0. case t_steps (start_of (layout_of aprog) as, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) + (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp' of + (s, ab) \ s = Suc (start_of (layout_of aprog) as) \ + dec_inv_1 (layout_of aprog) n e (as, am) (s, ab) ires \ + dec_inv_2 (layout_of aprog) n e (as, am) (s, ab) ires" + proof(erule_tac exE, simp, rule_tac x = "Suc 0" in exI, + simp add: t_steps.simps) + qed +next + assume "abc_fetch as aprog = Some (Dec n e)" + thus "0 < start_of (layout_of aprog) as" + apply(insert startof_not0[of "layout_of aprog" as], simp) + done +qed + +lemma dec_crsp_ex1: + "\crsp_l (layout_of aprog) (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e); + abc_lm_v am n = 0\ + \ \stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" +proof - + assume h1: "crsp_l (layout_of aprog) (as, am) tc ires" + "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0" + hence h2: "\ stp' > 0. (\ (s, l, r). + (s = Suc (start_of (layout_of aprog) as) \ + (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r)) ires)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')" + apply(insert crsp_next_state[of aprog as am tc ires n e], auto) + done + from h1 and h2 show + "\stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) + (t_steps tc (ci (layout_of aprog) (start_of + (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp) ires" + proof(erule_tac exE, case_tac "(t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), start_of + (layout_of aprog) as - Suc 0) stp')", simp) + fix stp' a b c + assume h3: "stp' > 0 \ a = Suc (start_of (layout_of aprog) as) \ + dec_inv_1 (layout_of aprog) n e (as, am) (a, b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0" + "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + thus "\stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" + proof(erule_tac conjE, simp) + assume "dec_inv_1 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" + "abc_lm_v am n = 0" + " t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + hence h4: "\stp. (\(s', l', r'). s' = + start_of (layout_of aprog) e \ + dec_inv_1 (layout_of aprog) n e (as, am) (s', l', r') ires) + (t_steps (start_of (layout_of aprog) as + 1, b, c) + (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp)" + apply(rule_tac dec_inv_stop1, auto) + done + from h3 and h4 show ?thesis + apply(erule_tac exE) + apply(rule_tac x = "stp' + stp" in exI, simp) + apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), + b, c) (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp)", + simp) + apply(rule_tac dec_inv_stop_cond1, auto) + done + qed + qed +qed + +lemma dec_crsp_ex2: + "\crsp_l (layout_of aprog) (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e); + 0 < abc_lm_v am n\ + \ \stp > 0. crsp_l (layout_of aprog) + (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" +proof - + assume h1: + "crsp_l (layout_of aprog) (as, am) tc ires" + "abc_fetch as aprog = Some (Dec n e)" + "abc_lm_v am n > 0" + hence h2: + "\ stp' > 0. (\ (s, l, r). (s = Suc (start_of (layout_of aprog) as) + \ (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) +(t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')" + apply(insert crsp_next_state[of aprog as am tc ires n e], auto) + done + from h1 and h2 show + "\stp >0. crsp_l (layout_of aprog) + (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires" + proof(erule_tac exE, + case_tac + "(t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')", simp) + fix stp' a b c + assume h3: "0 < stp' \ a = Suc (start_of (layout_of aprog) as) \ + dec_inv_2 (layout_of aprog) n e (as, am) (a, b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" + "abc_lm_v am n > 0" + "t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Dec n e), + start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + thus "?thesis" + proof(erule_tac conjE, simp) + assume + "dec_inv_2 (layout_of aprog) n e (as, am) + (Suc (start_of (layout_of aprog) as), b, c) ires" + "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n > 0" + "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp' + = (Suc (start_of (layout_of aprog) as), b, c)" + hence h4: + "\stp. (\(s', l', r'). s' = start_of (layout_of aprog) (Suc as) \ + dec_inv_2 (layout_of aprog) n e (as, am) (s', l', r') ires) + (t_steps (start_of (layout_of aprog) as + 1, b, c) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)" + apply(rule_tac dec_stop2, auto) + done + from h3 and h4 show ?thesis + apply(erule_tac exE) + apply(rule_tac x = "stp' + stp" in exI, simp) + apply(case_tac + "(t_steps (Suc (start_of (layout_of aprog) as), b, c) + (ci (layout_of aprog) (start_of (layout_of aprog) as) + (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)" + ,simp) + apply(rule_tac dec_inv_stop_cond2, auto) + done + qed + qed +qed + +lemma dec_crsp_ex_pre: + "\ly = layout_of aprog; crsp_l ly (as, am) tc ires; + abc_fetch as aprog = Some (Dec n e)\ + \ \stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e), + start_of ly as - Suc 0) stp) ires" +apply(auto simp: abc_step_l.simps intro: dec_crsp_ex2 dec_crsp_ex1) +done + +lemma dec_crsp_ex: + assumes layout: -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *} + "ly = layout_of aprog" + and dec: -- {* There is an @{text "Dec n e"} instruction at postion @{text "as"} of @{text "aprog"} *} + "abc_fetch as aprog = Some (Dec n e)" + and correspond: + -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM + configuration @{text "tc"} + *} + "crsp_l ly (as, am) tc ires" +shows + "\stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e), + start_of ly as - Suc 0) stp) ires" +proof - + from dec_crsp_ex_pre layout dec correspond show ?thesis by blast +qed + +(* +subsection {* Compilation of @{text "Goto n"}*} +*) + +lemma goto_fetch: + "fetch (ci (layout_of aprog) + (start_of (layout_of aprog) as) (Goto n)) (Suc 0) b + = (Nop, start_of (layout_of aprog) n)" +apply(auto simp: ci.simps fetch.simps nth_of.simps + split: block.splits) +done + +text {* + Correctness of complied @{text "Goto n"} + *} + +lemma goto_crsp_ex_pre: + "\ly = layout_of aprog; + crsp_l ly (as, am) tc ires; + abc_fetch as aprog = Some (Goto n)\ + \ \stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n), + start_of ly as - Suc 0) stp) ires" +apply(rule_tac x = 1 in exI) +apply(simp add: abc_step_l.simps t_steps.simps t_step.simps) +apply(case_tac tc, simp) +apply(subgoal_tac "a = start_of (layout_of aprog) as", auto) +apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp) +apply(auto simp: goto_fetch new_tape.simps crsp_l.simps) +apply(rule startof_not0) +done + +lemma goto_crsp_ex: + assumes layout: "ly = layout_of aprog" + and goto: "abc_fetch as aprog = Some (Goto n)" + and correspondence: "crsp_l ly (as, am) tc ires" + shows "\stp>0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) + (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n), + start_of ly as - Suc 0) stp) ires" +proof - + from goto_crsp_ex_pre and layout goto correspondence show "?thesis" by blast +qed + +subsection {* + The correctness of the compiler + *} + +declare abc_step_l.simps[simp del] + +lemma tm_crsp_ex: + "\ly = layout_of aprog; + crsp_l ly (as, am) tc ires; + as < length aprog; + abc_fetch as aprog = Some ins\ + \ \ n > 0. crsp_l ly (abc_step_l (as,am) (Some ins)) + (t_steps tc (ci (layout_of aprog) (start_of ly as) + (ins), (start_of ly as) - (Suc 0)) n) ires" +apply(case_tac "ins", simp) +apply(auto intro: inc_crsp_ex_pre dec_crsp_ex goto_crsp_ex) +done + +lemma start_of_pre: + "n < length aprog \ start_of (layout_of aprog) n + = start_of (layout_of (butlast aprog)) n" +apply(induct n, simp add: start_of.simps, simp) +apply(simp add: layout_of.simps start_of.simps) +apply(subgoal_tac "n < length aprog - Suc 0", simp) +apply(subgoal_tac "(aprog ! n) = (butlast aprog ! n)", simp) +proof - + fix n + assume h1: "Suc n < length aprog" + thus "aprog ! n = butlast aprog ! n" + apply(case_tac "length aprog", simp, simp) + apply(insert nth_append[of "butlast aprog" "[last aprog]" n]) + apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog") + apply(simp split: if_splits) + apply(rule append_butlast_last_id, case_tac aprog, simp, simp) + done +next + fix n + assume "Suc n < length aprog" + thus "n < length aprog - Suc 0" + apply(case_tac aprog, simp, simp) + done +qed + +lemma zip_eq: "xs = ys \ zip xs zs = zip ys zs" +by simp + +lemma tpairs_of_append_iff: "length aprog = Suc n \ + tpairs_of aprog = tpairs_of (butlast aprog) @ + [(start_of (layout_of aprog) n, aprog ! n)]" +apply(simp add: tpairs_of.simps) +apply(insert zip_append[of "map (start_of (layout_of aprog)) [0..(n, tm). abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0..aprog. x = length aprog \ + list_all (\(n, tm). abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0..(n, tm). abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0.. xs a. aprog = xs @ [a]" + using h + apply(rule_tac x = "butlast aprog" in exI, + rule_tac x = "last aprog" in exI) + apply(case_tac "aprog = []", simp, simp) + done + from this obtain xs where "\ a. aprog = xs @ [a]" .. + from this obtain a where g3: "aprog = xs @ [a]" .. + from g1 and g2 and g3 show "list_all (\(n, tm). + abacus.t_ncorrect (ci layout n tm)) + (zip (map (start_of layout) [0..ly = layout_of aprog; tprog = tm_of aprog; + crsp_l ly (as, am) tc ires; length aprog \ as\ + \ abc_step_l (as, am) (abc_fetch as aprog) = (as, am)" +apply(simp add: abc_fetch.simps abc_step_l.simps) +done + +lemma tm_merge_ex: + "\crsp_l (layout_of aprog) (as, am) tc ires; + as < length aprog; + abc_fetch as aprog = Some a; + abc2t_correct aprog; + crsp_l (layout_of aprog) (abc_step_l (as, am) (Some a)) + (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) + a, start_of (layout_of aprog) as - Suc 0) n) ires; + n > 0\ + \ \stp > 0. crsp_l (layout_of aprog) (abc_step_l (as, am) + (Some a)) (t_steps tc (tm_of aprog, 0) stp) ires" +apply(case_tac "(t_steps tc (ci (layout_of aprog) + (start_of (layout_of aprog) as) a, + start_of (layout_of aprog) as - Suc 0) n)", simp) +apply(case_tac "(abc_step_l (as, am) (Some a))", simp) +proof - + fix aa b c aaa ba + assume h: + "crsp_l (layout_of aprog) (as, am) tc ires" + "as < length aprog" + "abc_fetch as aprog = Some a" + "crsp_l (layout_of aprog) (aaa, ba) (aa, b, c) ires" + "abc2t_correct aprog" + "n>0" + "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) a, + start_of (layout_of aprog) as - Suc 0) n = (aa, b, c)" + "abc_step_l (as, am) (Some a) = (aaa, ba)" + hence "aa = start_of (layout_of aprog) aaa" + apply(simp add: crsp_l.simps) + done + from this and h show + "\stp > 0. crsp_l (layout_of aprog) (aaa, ba) + (t_steps tc (tm_of aprog, 0) stp) ires" + apply(insert tms_out_ex[of "layout_of aprog" aprog + "tm_of aprog" as am tc ires a n aa b c aaa ba], auto) + done +qed + +lemma crsp_inside: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly (as, am) tc ires; + as < length aprog\ \ + (\ stp > 0. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) + (t_steps tc (tprog, 0) stp) ires)" +apply(case_tac "abc_fetch as aprog", simp add: abc_fetch.simps) +proof - + fix a + assume "ly = layout_of aprog" + "tprog = tm_of aprog" + "crsp_l ly (as, am) tc ires" + "as < length aprog" + "abc_fetch as aprog = Some a" + thus "\stp > 0. crsp_l ly (abc_step_l (as, am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires" + proof(insert tm_crsp_ex[of ly aprog as am tc ires a], + auto intro: tm_merge_ex) + qed +qed + +lemma crsp_outside: + "\ly = layout_of aprog; tprog = tm_of aprog; + crsp_l ly (as, am) tc ires; as \ length aprog\ + \ (\ stp. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) + (t_steps tc (tprog, 0) stp) ires)" +apply(subgoal_tac "abc_step_l (as, am) (abc_fetch as aprog) + = (as, am)", simp) +apply(rule_tac x = 0 in exI, simp add: t_steps.simps) +apply(rule as_out, simp+) +done + +text {* + Single-step correntess of the compiler. +*} +lemma astep_crsp_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly (as, am) tc ires\ + \ (\ stp. crsp_l ly (abc_step_l (as,am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" +apply(case_tac "as < length aprog") +apply(drule_tac crsp_inside, auto) +apply(rule_tac crsp_outside, simp+) +done + +text {* + Single-step correntess of the compiler. +*} +lemma astep_crsp_pre1: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly (as, am) tc ires\ + \ (\ stp. crsp_l ly (abc_step_l (as,am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" +apply(case_tac "as < length aprog") +apply(drule_tac crsp_inside, auto) +apply(rule_tac crsp_outside, simp+) +done + +lemma astep_crsp: + assumes layout: + -- {* There is a Abacus program @{text "aprog"} with layout @{text "ly"} *} + "ly = layout_of aprog" + and compiled: + -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *} + "tprog = tm_of aprog" + and corresponds: + -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM configuration + @{text "tc"} *} + "crsp_l ly (as, am) tc ires" + -- {* One step execution of @{text "aprog"} can be simulated by multi-step execution + of @{text "tprog"} *} + shows "(\ stp. crsp_l ly (abc_step_l (as,am) + (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" +proof - + from astep_crsp_pre1 [OF layout compiled corresponds] show ?thesis . +qed + +lemma steps_crsp_pre: + "\ly = layout_of aprog; tprog = tm_of aprog; + crsp_l ly ac tc ires; ac' = abc_steps_l ac aprog n\ \ + (\ n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)" +apply(induct n arbitrary: ac' ac tc, simp add: abc_steps_l.simps) +apply(rule_tac x = 0 in exI) +apply(case_tac ac, simp add: abc_steps_l.simps t_steps.simps) +apply(case_tac ac, simp add: abc_steps_l.simps) +apply(subgoal_tac + "(\ stp. crsp_l ly (abc_step_l (a, b) + (abc_fetch a aprog)) (t_steps tc (tprog, 0) stp) ires)") +apply(erule exE) +apply(subgoal_tac + "\n'. crsp_l (layout_of aprog) + (abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) aprog n) + (t_steps ((t_steps tc (tprog, 0) stp)) (tm_of aprog, 0) n') ires") +apply(erule exE) +apply(subgoal_tac + "t_steps (t_steps tc (tprog, 0) stp) (tm_of aprog, 0) n' = + t_steps tc (tprog, 0) (stp + n')") +apply(rule_tac x = "stp + n'" in exI, simp) +apply(auto intro: astep_crsp simp: t_step_add) +done + +text {* + Multi-step correctess of the compiler. +*} + +lemma steps_crsp: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *} + "ly = layout_of aprog" + and compiled: + -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *} + "tprog = tm_of aprog" + and correspond: + -- {* Abacus configuration @{text "ac"} is in correspondence with TM configuration @{text "tc"} *} + "crsp_l ly ac tc ires" + and execution: + -- {* @{text "ac'"} is the configuration obtained from @{text "n"}-step execution + of @{text "aprog"} starting from configuration @{text "ac"} *} + "ac' = abc_steps_l ac aprog n" + -- {* There exists steps @{text "n'"} steps, after these steps of execution, + the Turing configuration such obtained is in correspondence with @{text "ac'"} *} + shows "(\ n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)" +proof - + from steps_crsp_pre [OF layout compiled correspond execution] show ?thesis . +qed + +subsection {* The Mop-up machine *} + +fun mop_bef :: "nat \ tprog" + where + "mop_bef 0 = []" | + "mop_bef (Suc n) = mop_bef n @ + [(R, 2*n + 3), (W0, 2*n + 2), (R, 2*n + 1), (W1, 2*n + 2)]" + +definition mp_up :: "tprog" + where + "mp_up \ [(R, 2), (R, 1), (L, 5), (W0, 3), (R, 4), (W0, 3), + (R, 2), (W0, 3), (L, 5), (L, 6), (R, 0), (L, 6)]" + +fun tMp :: "nat \ nat \ tprog" + where + "tMp n off = tshift (mop_bef n @ tshift mp_up (2*n)) off" + +declare mp_up_def[simp del] tMp.simps[simp del] mop_bef.simps[simp del] +(**********Begin: equiv among aba and turing***********) + +lemma tm_append_step: + "\t_ncorrect tp1; t_step tc (tp1, 0) = (s, l, r); s \ 0\ + \ t_step tc (tp1 @ tp2, 0) = (s, l, r)" +apply(simp add: t_step.simps) +apply(case_tac tc, simp) +apply(case_tac + "(fetch tp1 a (case c of [] \ Bk | + Bk # xs \ Bk | Oc # xs \ Oc))", simp) +apply(case_tac a, simp add: fetch.simps) +apply(simp add: fetch.simps) +apply(case_tac c, simp) +apply(case_tac [!] "ab::block") +apply(auto simp: nth_of.simps nth_append t_ncorrect.simps + split: if_splits) +done + +lemma state0_ind: "t_steps (0, l, r) (tp, 0) stp = (0, l, r)" +apply(induct stp, simp add: t_steps.simps) +apply(simp add: t_steps.simps t_step.simps fetch.simps new_tape.simps) +done + +lemma tm_append_steps: + "\t_ncorrect tp1; t_steps tc (tp1, 0) stp = (s, l ,r); s \ 0\ + \ t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" +apply(induct stp arbitrary: tc s l r) +apply(case_tac tc, simp) +apply(simp add: t_steps.simps) +proof - + fix stp tc s l r + assume h1: "\tc s l r. \t_ncorrect tp1; t_steps tc (tp1, 0) stp = + (s, l, r); s \ 0\ \ t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" + and h2: "t_steps tc (tp1, 0) (Suc stp) = (s, l, r)" "s \ 0" + "t_ncorrect tp1" + thus "t_steps tc (tp1 @ tp2, 0) (Suc stp) = (s, l, r)" + apply(simp add: t_steps.simps) + apply(case_tac "(t_step tc (tp1, 0))", simp) + proof- + fix a b c + assume g1: "\tc s l r. \t_steps tc (tp1, 0) stp = (s, l, r); + 0 < s\ \ t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" + and g2: "t_step tc (tp1, 0) = (a, b, c)" + "t_steps (a, b, c) (tp1, 0) stp = (s, l, r)" + "0 < s" + "t_ncorrect tp1" + hence g3: "a > 0" + apply(case_tac "a::nat", auto simp: t_steps.simps) + apply(simp add: state0_ind) + done + from g1 and g2 and this have g4: + "(t_step tc (tp1 @ tp2, 0)) = (a, b, c)" + apply(rule_tac tm_append_step, simp, simp, simp) + done + from g1 and g2 and g3 and g4 show + "t_steps (t_step tc (tp1 @ tp2, 0)) (tp1 @ tp2, 0) stp + = (s, l, r)" + apply(simp) + done + qed +qed + +lemma shift_fetch: + "\n < length tp; + (tp:: (taction \ nat) list) ! n = (aa, ba); + ba \ 0\ + \ (tshift tp (length tp div 2)) ! n = + (aa , ba + length tp div 2)" +apply(simp add: tshift.simps) +done + +lemma tshift_length_equal: "length (tshift tp q) = length tp" +apply(auto simp: tshift.simps) +done + +thm nth_of.simps + +lemma [simp]: "t_ncorrect tp \ 2 * (length tp div 2) = length tp" +apply(auto simp: t_ncorrect.simps) +done + +lemma tm_append_step_equal': + "\t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\ \ + (\ (s, l, r). ((\ (s', l', r'). + (s'\ 0 \ (s = s' + off \ l = l' \ r = r'))) + (t_step (a, b, c) (tp', 0)))) + (t_step (a + off, b, c) (tp @ tshift tp' off, 0))" +apply(simp add: t_step.simps) +apply(case_tac a, simp add: fetch.simps) +apply(case_tac +"(fetch tp' a (case c of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", + simp) +apply(case_tac +"(fetch (tp @ tshift tp' (length tp div 2)) + (Suc (nat + length tp div 2)) + (case c of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", + simp) +apply(case_tac "(new_tape aa (b, c))", + case_tac "(new_tape aaa (b, c))", simp, + rule impI, simp add: fetch.simps split: block.splits option.splits) +apply (auto simp: nth_of.simps t_ncorrect.simps + nth_append tshift_length_equal tshift.simps split: if_splits) +done + + +lemma tm_append_step_equal: + "\t_ncorrect tp; t_ncorrect tp'; off = length tp div 2; + t_step (a, b, c) (tp', 0) = (aa, ab, bb); aa \ 0\ + \ t_step (a + length tp div 2, b, c) + (tp @ tshift tp' (length tp div 2), 0) + = (aa + length tp div 2, ab, bb)" +apply(insert tm_append_step_equal'[of tp tp' off a b c], simp) +apply(case_tac "(t_step (a + length tp div 2, b, c) + (tp @ tshift tp' (length tp div 2), 0))", simp) +done + +lemma tm_append_steps_equal: + "\t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\ \ + (\ (s, l, r). ((\ (s', l', r'). ((s'\ 0 \ s = s' + off \ l = l' + \ r = r'))) (t_steps (a, b, c) (tp', 0) stp))) + (t_steps (a + off, b, c) (tp @ tshift tp' off, 0) stp)" +apply(induct stp arbitrary : a b c, simp add: t_steps.simps) +apply(simp add: t_steps.simps) +apply(case_tac "(t_step (a, b, c) (tp', 0))", simp) +apply(case_tac "aa = 0", simp add: state0_ind) +apply(subgoal_tac "(t_step (a + length tp div 2, b, c) + (tp @ tshift tp' (length tp div 2), 0)) + = (aa + length tp div 2, ba, ca)", simp) +apply(rule tm_append_step_equal, auto) +done + +(*********Begin: mop_up***************) +type_synonym mopup_type = "t_conf \ nat list \ nat \ block list \ bool" + +fun mopup_stop :: "mopup_type" + where + "mopup_stop (s, l, r) lm n ires= + (\ ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r = @ Bk\<^bsup>rn\<^esup>)" + +fun mopup_bef_erase_a :: "mopup_type" + where + "mopup_bef_erase_a (s, l, r) lm n ires= + (\ ln m rn. l = Bk \<^bsup>ln\<^esup> @ Bk # Bk # ires \ + r = Oc\<^bsup>m \<^esup>@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<^bsup>rn\<^esup>)" + +fun mopup_bef_erase_b :: "mopup_type" + where + "mopup_bef_erase_b (s, l, r) lm n ires = + (\ ln m rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r = Bk # Oc\<^bsup>m\<^esup> @ Bk # + <(drop (s div 2) lm)> @ Bk\<^bsup>rn\<^esup>)" + + +fun mopup_jump_over1 :: "mopup_type" + where + "mopup_jump_over1 (s, l, r) lm n ires = + (\ ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \ + l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ + (r = Oc\<^bsup>m2\<^esup> @ Bk # <(drop (Suc n) lm)> @ Bk\<^bsup>rn \<^esup>\ + (r = Oc\<^bsup>m2\<^esup> \ (drop (Suc n) lm) = [])))" + +fun mopup_aft_erase_a :: "mopup_type" + where + "mopup_aft_erase_a (s, l, r) lm n ires = + (\ lnl lnr rn (ml::nat list) m. + m = Suc (abc_lm_v lm n) \ l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ + (r = @ Bk\<^bsup>rn\<^esup>))" + +fun mopup_aft_erase_b :: "mopup_type" + where + "mopup_aft_erase_b (s, l, r) lm n ires= + (\ lnl lnr rn (ml::nat list) m. + m = Suc (abc_lm_v lm n) \ + l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ + (r = Bk # @ Bk\<^bsup>rn \<^esup>\ + r = Bk # Bk # @ Bk\<^bsup>rn\<^esup>))" + +fun mopup_aft_erase_c :: "mopup_type" + where + "mopup_aft_erase_c (s, l, r) lm n ires = + (\ lnl lnr rn (ml::nat list) m. + m = Suc (abc_lm_v lm n) \ + l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ + (r = @ Bk\<^bsup>rn \<^esup>\ r = Bk # @ Bk\<^bsup>rn\<^esup>))" + +fun mopup_left_moving :: "mopup_type" + where + "mopup_left_moving (s, l, r) lm n ires = + (\ lnl lnr rn m. + m = Suc (abc_lm_v lm n) \ + ((l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ r = Bk\<^bsup>rn\<^esup>) \ + (l = Oc\<^bsup>m - 1\<^esup> @ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \ r = Oc # Bk\<^bsup>rn\<^esup>)))" + +fun mopup_jump_over2 :: "mopup_type" + where + "mopup_jump_over2 (s, l, r) lm n ires = + (\ ln rn m1 m2. + m1 + m2 = Suc (abc_lm_v lm n) + \ r \ [] + \ (hd r = Oc \ (l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \ r = Oc\<^bsup>m2\<^esup> @ Bk\<^bsup>rn\<^esup>)) + \ (hd r = Bk \ (l = Bk\<^bsup>ln\<^esup> @ Bk # ires \ r = Bk # Oc\<^bsup>m1 + m2\<^esup> @ Bk\<^bsup>rn\<^esup>)))" + + +fun mopup_inv :: "mopup_type" + where + "mopup_inv (s, l, r) lm n ires = + (if s = 0 then mopup_stop (s, l, r) lm n ires + else if s \ 2*n then + if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires + else mopup_bef_erase_b (s, l, r) lm n ires + else if s = 2*n + 1 then + mopup_jump_over1 (s, l, r) lm n ires + else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires + else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires + else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires + else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires + else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires + else False)" + +declare + mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del] + mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del] + mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del] + mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del] + mopup_stop.simps[simp del] + +lemma mopup_fetch_0[simp]: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) 0 b) = (Nop, 0)" +by(simp add: fetch.simps) + +lemma mop_bef_length[simp]: "length (mop_bef n) = 4 * n" +apply(induct n, simp add: mop_bef.simps, simp add: mop_bef.simps) +done + +thm findnth_nth +lemma mop_bef_nth: + "\q < n; x < 4\ \ mop_bef n ! (4 * q + x) = + mop_bef (Suc q) ! ((4 * q) + x)" +apply(induct n, simp) +apply(case_tac "q < n", simp add: mop_bef.simps, auto) +apply(simp add: nth_append) +apply(subgoal_tac "q = n", simp) +apply(arith) +done + +lemma fetch_bef_erase_a_o[simp]: + "\0 < s; s \ 2 * n; s mod 2 = Suc 0\ + \ (fetch (mop_bef n @ tshift mp_up (2 * n)) s Oc) = (W0, s + 1)" +apply(subgoal_tac "\ q. s = 2*q + 1", auto) +apply(subgoal_tac "length (mop_bef n) = 4*n") +apply(auto simp: fetch.simps nth_of.simps nth_append) +apply(subgoal_tac "mop_bef n ! (4 * q + 1) = + mop_bef (Suc q) ! ((4 * q) + 1)", + simp add: mop_bef.simps nth_append) +apply(rule mop_bef_nth, auto) +done + +lemma fetch_bef_erase_a_b[simp]: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = Suc 0\ + \ (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s + 2)" +apply(subgoal_tac "\ q. s = 2*q + 1", auto) +apply(subgoal_tac "length (mop_bef n) = 4*n") +apply(auto simp: fetch.simps nth_of.simps nth_append) +apply(subgoal_tac "mop_bef n ! (4 * q + 0) = + mop_bef (Suc q) ! ((4 * q + 0))", + simp add: mop_bef.simps nth_append) +apply(rule mop_bef_nth, auto) +done + +lemma fetch_bef_erase_b_b: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = 0\ \ + (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s - 1)" +apply(subgoal_tac "\ q. s = 2 * q", auto) +apply(case_tac qa, simp, simp) +apply(auto simp: fetch.simps nth_of.simps nth_append) +apply(subgoal_tac "mop_bef n ! (4 * nat + 2) = + mop_bef (Suc nat) ! ((4 * nat) + 2)", + simp add: mop_bef.simps nth_append) +apply(rule mop_bef_nth, auto) +done + +lemma fetch_jump_over1_o: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Oc + = (R, Suc (2 * n))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def nth_append + tshift.simps) +done + +lemma fetch_jump_over1_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Bk + = (R, Suc (Suc (2 * n)))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def + nth_append tshift.simps) +done + +lemma fetch_aft_erase_a_o: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Oc + = (W0, Suc (2 * n + 2))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def + nth_append tshift.simps) +done + +lemma fetch_aft_erase_a_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Bk + = (L, Suc (2 * n + 4))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(auto simp: fetch.simps nth_of.simps mp_up_def + nth_append tshift.simps) +done + +lemma fetch_aft_erase_b_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (2*n + 3) Bk + = (R, Suc (2 * n + 3))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_aft_erase_c_o: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Oc + = (W0, Suc (2 * n + 2))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_aft_erase_c_b: + "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Bk + = (R, Suc (2 * n + 1))" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_left_moving_o: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Oc) + = (L, 2*n + 6)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_left_moving_b: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Bk) + = (L, 2*n + 5)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_jump_over2_b: + "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Bk) + = (R, 0)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemma fetch_jump_over2_o: +"(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Oc) + = (L, 2*n + 6)" +apply(subgoal_tac "length (mop_bef n) = 4 * n") +apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) +apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) +done + +lemmas mopupfetchs = +fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b +fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o +fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o +fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b +fetch_jump_over2_b fetch_jump_over2_o + +lemma [simp]: +"\n < length lm; 0 < s; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Oc # xs) lm n ires; + Suc s \ 2 * n\ \ + mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps ) +apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI) +apply(case_tac m, simp, simp) +done + +lemma mopup_false1: + "\0 < s; s \ 2 * n; s mod 2 = Suc 0; \ Suc s \ 2 * n\ + \ RR" +apply(arith) +done + +lemma [simp]: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\ + \ (Suc s \ 2 * n \ mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires) \ + (\ Suc s \ 2 * n \ mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) " +apply(auto elim: mopup_false1) +done + +lemma drop_abc_lm_v_simp[simp]: + "n < length lm \ drop n lm = abc_lm_v lm n # drop (Suc n) lm" +apply(auto simp: abc_lm_v.simps) +apply(drule drop_Suc_conv_tl, simp) +done +lemma [simp]: "(\rna. Bk\<^bsup>rn\<^esup> = Bk # Bk\<^bsup>rna\<^esup>) \ Bk\<^bsup>rn\<^esup> = []" +apply(case_tac rn, simp, auto) +done + +lemma [simp]: "\lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup>" +apply(rule_tac x = "Suc ln" in exI, auto) +done + +lemma mopup_bef_erase_a_2_jump_over[simp]: + "\n < length lm; 0 < s; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Bk # xs) lm n ires; Suc s = 2 * n\ +\ mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps) +apply(case_tac m, simp) +apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI, + simp add: tape_of_nl_abv) +apply(case_tac "drop (Suc n) lm", auto simp: tape_of_nat_list.simps ) +done + +lemma Suc_Suc_div: "\0 < s; s mod 2 = Suc 0; Suc (Suc s) \ 2 * n\ + \ (Suc (Suc (s div 2))) \ n" +apply(arith) +done + +lemma mopup_bef_erase_a_2_a[simp]: + "\n < length lm; 0 < s; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Bk # xs) lm n ires; + Suc (Suc s) \ 2 * n\ \ + mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps ) +apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \ []") +apply(case_tac m, simp) +apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI, + rule_tac x = rn in exI, simp, simp) +apply(subgoal_tac "(Suc (Suc (s div 2))) \ n", simp, + rule_tac Suc_Suc_div, auto) +done + +lemma mopup_false2: + "\n < length lm; 0 < s; s \ 2 * n; + s mod 2 = Suc 0; Suc s \ 2 * n; + \ Suc (Suc s) \ 2 * n\ \ RR" +apply(arith) +done + +lemma [simp]: + "\n < length lm; 0 < s; s \ 2 * n; + s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, Bk # xs) lm n ires; + r = Bk # xs\ + \ (Suc s = 2 * n \ + mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires) \ + (Suc s \ 2 * n \ + (Suc (Suc s) \ 2 * n \ + mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires) \ + (\ Suc (Suc s) \ 2 * n \ + mopup_aft_erase_a (Suc (Suc s), Bk # l, xs) lm n ires))" +apply(auto elim: mopup_false2) +done + +lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \ + mopup_bef_erase_a (s, l, [Bk]) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps) +done + +lemma [simp]: + "\n < length lm; 0 < s; s \ 2 * n; s mod 2 = Suc 0; + mopup_bef_erase_a (s, l, []) lm n ires; r = []\ + \ (Suc s = 2 * n \ + mopup_jump_over1 (Suc (2 * n), Bk # l, []) lm n ires) \ + (Suc s \ 2 * n \ + (Suc (Suc s) \ 2 * n \ + mopup_bef_erase_a (Suc (Suc s), Bk # l, []) lm n ires) \ + (\ Suc (Suc s) \ 2 * n \ + mopup_aft_erase_a (Suc (Suc s), Bk # l, []) lm n ires))" +apply(auto) +done + +lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \ l \ []" +apply(auto simp: mopup_bef_erase_b.simps) +done + +lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False" +apply(auto simp: mopup_bef_erase_b.simps ) +done + +lemma [simp]: "\0 < s; s \ 2 *n; s mod 2 \ Suc 0\ \ + (s - Suc 0) mod 2 = Suc 0" +apply(arith) +done + +lemma [simp]: "\0 < s; s \ 2 *n; s mod 2 \ Suc 0\ \ + s - Suc 0 \ 2 * n" +apply(simp) +done + +lemma [simp]: "\0 < s; s \ 2 *n; s mod 2 \ Suc 0\ \ \ s \ Suc 0" +apply(arith) +done + +lemma [simp]: "\n < length lm; 0 < s; s \ 2 * n; + s mod 2 \ Suc 0; + mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\ + \ mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires" +apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) +done + +lemma [simp]: "\mopup_bef_erase_b (s, l, []) lm n ires\ \ + mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires" +apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) +done + +lemma [simp]: + "\n < length lm; + mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires; + r = Oc # xs\ + \ mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires" +apply(auto simp: mopup_jump_over1.simps) +apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, + rule_tac x = "m2 - 1" in exI) +apply(case_tac "m2", simp, simp, rule_tac x = rn in exI, simp) +apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, + rule_tac x = "m2 - 1" in exI) +apply(case_tac m2, simp, simp) +done + +lemma mopup_jump_over1_2_aft_erase_a[simp]: + "\n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\ + \ mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" +apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps) +apply(erule_tac exE)+ +apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI) +apply(case_tac m2, simp) +apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI, + simp) +apply(simp) +done + +lemma [simp]: + "\n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\ \ + mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" +apply(rule mopup_jump_over1_2_aft_erase_a, simp) +apply(auto simp: mopup_jump_over1.simps) +apply(rule_tac x = ln in exI, rule_tac x = m1 in exI, + rule_tac x = m2 in exI, simp add: ) +apply(rule_tac x = 0 in exI, auto) +done + +lemma [simp]: + "\n < length lm; + mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\ + \ mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" +apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps ) +apply(case_tac ml, simp, case_tac rn, simp, simp) +apply(case_tac list, auto simp: tape_of_nl_abv + tape_of_nat_list.simps ) +apply(case_tac a, simp, rule_tac x = rn in exI, + rule_tac x = "[]" in exI, + simp add: tape_of_nat_list.simps, simp) +apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, + simp add: tape_of_nat_list.simps ) +apply(case_tac a, simp, rule_tac x = rn in exI, + rule_tac x = "aa # lista" in exI, simp, simp) +apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, + simp add: tape_of_nat_list.simps ) +done + +lemma [simp]: + "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \ l \ []" +apply(auto simp: mopup_aft_erase_a.simps) +done + +lemma [simp]: + "\n < length lm; + mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\ + \ mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires" +apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) +apply(erule exE)+ +apply(case_tac lnr, simp) +apply(rule_tac x = lnl in exI, simp, rule_tac x = rn in exI, simp) +apply(subgoal_tac "ml = []", simp) +apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, auto) +apply(subgoal_tac "ml = []", auto) +apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp) +done + +lemma [simp]: + "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \ l \ []" +apply(simp only: mopup_aft_erase_a.simps) +apply(erule exE)+ +apply(auto) +done + +lemma [simp]: + "\n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\ + \ mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires" +apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) +apply(erule exE)+ +apply(subgoal_tac "ml = [] \ rn = 0", erule conjE, erule conjE, simp) +apply(case_tac lnr, simp, rule_tac x = lnl in exI, simp, + rule_tac x = 0 in exI, simp) +apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, + rule_tac x = "Suc 0" in exI, simp) +apply(case_tac ml, simp, case_tac rn, simp, simp) +apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) +done + +lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False" +apply(auto simp: mopup_aft_erase_b.simps ) +done + +lemma [simp]: + "\n < length lm; + mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\ + \ mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" +apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps ) +apply(case_tac ml, simp, case_tac rn, simp, simp, simp) +apply(case_tac list, auto simp: tape_of_nl_abv + tape_of_nat_list.simps tape_of_nat_abv ) +apply(case_tac a, rule_tac x = rn in exI, + rule_tac x = "[]" in exI, simp add: tape_of_nat_list.simps) +apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, + simp add: tape_of_nat_list.simps ) +apply(case_tac a, simp, rule_tac x = rn in exI, + rule_tac x = "aa # lista" in exI, simp) +apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, + simp add: tape_of_nat_list.simps ) +done + +lemma mopup_aft_erase_c_aft_erase_a[simp]: + "\n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\ + \ mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" +apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps ) +apply(erule_tac exE)+ +apply(erule conjE, erule conjE, erule disjE) +apply(subgoal_tac "ml = []", simp, case_tac rn, + simp, simp, rule conjI) +apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) +apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp) +apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, simp, + rule conjI) +apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) +apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp) +done + +lemma [simp]: + "\n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\ + \ mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" +apply(rule mopup_aft_erase_c_aft_erase_a, simp) +apply(simp only: mopup_aft_erase_c.simps) +apply(erule exE)+ +apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: ) +apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp) +done + +lemma mopup_aft_erase_b_2_aft_erase_c[simp]: + "\n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\ + \ mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires" +apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps) +apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) +apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) +done + +lemma [simp]: + "\n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\ + \ mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires" +apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp) +apply(simp add: mopup_aft_erase_b.simps) +done + +lemma [simp]: + "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \ l \ []" +apply(auto simp: mopup_left_moving.simps) +done + +lemma [simp]: + "\n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\ + \ mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" +apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps) +apply(erule_tac exE)+ +apply(erule conjE, erule disjE, erule conjE) +apply(case_tac rn, simp, simp add: ) +apply(case_tac "hd l", simp add: ) +apply(case_tac "abc_lm_v lm n", simp) +apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI, + rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI) +apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp) +apply(case_tac "abc_lm_v lm n", simp) +apply(case_tac lnl, simp, simp) +apply(rule_tac x = lnl in exI, rule_tac x = rn in exI) +apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp) +done + +lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \ l \ []" +apply(auto simp: mopup_left_moving.simps) +done + +lemma [simp]: + "\n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\ + \ mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires" +apply(simp only: mopup_left_moving.simps) +apply(erule exE)+ +apply(case_tac lnr, simp) +apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, + rule_tac x = rn in exI, simp, simp) +apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp) +done + +lemma [simp]: +"\n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\ + \ mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires" +apply(simp only: mopup_left_moving.simps) +apply(erule exE)+ +apply(case_tac lnr, simp) +apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, + rule_tac x = 0 in exI, simp, auto) +done + +lemma [simp]: + "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \ l \ []" +apply(auto simp: mopup_jump_over2.simps ) +done + +lemma [intro]: "\lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup> @ [Bk]" +apply(simp only: exp_ind[THEN sym], auto) +done + +lemma [simp]: +"\n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\ + \ mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" +apply(simp only: mopup_jump_over2.simps) +apply(erule_tac exE)+ +apply(simp add: , erule conjE, erule_tac conjE) +apply(case_tac m1, simp) +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, + rule_tac x = 0 in exI, simp) +apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp) +apply(rule_tac x = ln in exI, rule_tac x = rn in exI, + rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp) +done + +lemma [simp]: "\rna. Oc # Oc\<^bsup>a\<^esup> @ Bk\<^bsup>rn\<^esup> = @ Bk\<^bsup>rna\<^esup>" +apply(case_tac a, auto simp: tape_of_nat_abv ) +done + +lemma [simp]: + "\n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\ + \ mopup_stop (0, Bk # l, xs) lm n ires" +apply(auto simp: mopup_jump_over2.simps mopup_stop.simps) +done + +lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False" +apply(simp only: mopup_jump_over2.simps, simp) +done + +lemma mopup_inv_step: + "\n < length lm; mopup_inv (s, l, r) lm n ires\ + \ mopup_inv (t_step (s, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0)) lm n ires" +apply(auto split:if_splits simp add:t_step.simps, + tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *}) +apply(simp_all add: mopupfetchs new_tape.simps) +done + +declare mopup_inv.simps[simp del] + +lemma mopup_inv_steps: +"\n < length lm; mopup_inv (s, l, r) lm n ires\ \ + mopup_inv (t_steps (s, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) stp) lm n ires" +apply(induct stp, simp add: t_steps.simps) +apply(simp add: t_steps_ind) +apply(case_tac "(t_steps (s, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp) +apply(rule_tac mopup_inv_step, simp, simp) +done + +lemma [simp]: + "\n < length lm; Suc 0 \ n\ \ + mopup_bef_erase_a (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) lm n ires" +apply(auto simp: mopup_bef_erase_a.simps abc_lm_v.simps) +apply(case_tac lm, simp, case_tac list, simp, simp) +apply(rule_tac x = "Suc a" in exI, rule_tac x = rn in exI, simp) +done + +lemma [simp]: + "lm \ [] \ mopup_jump_over1 (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) lm 0 ires" +apply(auto simp: mopup_jump_over1.simps) +apply(rule_tac x = ln in exI, rule_tac x = 0 in exI, simp add: ) +apply(case_tac lm, simp, simp add: abc_lm_v.simps) +apply(case_tac rn, simp) +apply(case_tac list, rule_tac disjI2, + simp add: tape_of_nl_abv tape_of_nat_list.simps) +apply(rule_tac disjI1, + simp add: tape_of_nl_abv tape_of_nat_list.simps ) +apply(rule_tac disjI1, case_tac list, + simp add: tape_of_nl_abv tape_of_nat_list.simps, + rule_tac x = nat in exI, simp) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps ) +done + +lemma mopup_init: + "\n < length lm; crsp_l ly (as, lm) (ac, l, r) ires\ \ + mopup_inv (Suc 0, l, r) lm n ires" +apply(auto simp: crsp_l.simps mopup_inv.simps) +apply(case_tac n, simp, auto simp: mopup_bef_erase_a.simps ) +apply(rule_tac x = "Suc (hd lm)" in exI, rule_tac x = rn in exI, simp) +apply(case_tac lm, simp, case_tac list, simp, case_tac lista, simp add: abc_lm_v.simps) +apply(simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(simp add: mopup_jump_over1.simps) +apply(rule_tac x = 0 in exI, rule_tac x = 0 in exI, auto) +apply(case_tac [!] n, simp_all) +apply(case_tac [!] lm, simp, case_tac list, simp) +apply(case_tac rn, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(erule_tac x = nat in allE, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(simp add: abc_lm_v.simps, auto) +apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) +apply(erule_tac x = rn in allE, simp_all) +done + +fun abc_mopup_stage1 :: "t_conf \ nat \ nat" + where + "abc_mopup_stage1 (s, l, r) n = + (if s > 0 \ s \ 2*n then 6 + else if s = 2*n + 1 then 4 + else if s \ 2*n + 2 \ s \ 2*n + 4 then 3 + else if s = 2*n + 5 then 2 + else if s = 2*n + 6 then 1 + else 0)" + +fun abc_mopup_stage2 :: "t_conf \ nat \ nat" + where + "abc_mopup_stage2 (s, l, r) n = + (if s > 0 \ s \ 2*n then length r + else if s = 2*n + 1 then length r + else if s = 2*n + 5 then length l + else if s = 2*n + 6 then length l + else if s \ 2*n + 2 \ s \ 2*n + 4 then length r + else 0)" + +fun abc_mopup_stage3 :: "t_conf \ nat \ nat" + where + "abc_mopup_stage3 (s, l, r) n = + (if s > 0 \ s \ 2*n then + if hd r = Bk then 0 + else 1 + else if s = 2*n + 2 then 1 + else if s = 2*n + 3 then 0 + else if s = 2*n + 4 then 2 + else 0)" + +fun abc_mopup_measure :: "(t_conf \ nat) \ (nat \ nat \ nat)" + where + "abc_mopup_measure (c, n) = + (abc_mopup_stage1 c n, abc_mopup_stage2 c n, + abc_mopup_stage3 c n)" + +definition abc_mopup_LE :: + "(((nat \ block list \ block list) \ nat) \ + ((nat \ block list \ block list) \ nat)) set" + where + "abc_mopup_LE \ (inv_image lex_triple abc_mopup_measure)" + +lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE" +by(auto intro:wf_inv_image wf_lex_triple simp:abc_mopup_LE_def) + +lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False" +apply(auto simp: mopup_bef_erase_a.simps) +done + +lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False" +apply(auto simp: mopup_bef_erase_b.simps) +done + +lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False" +apply(auto simp: mopup_aft_erase_b.simps) +done + +lemma mopup_halt_pre: + "\n < length lm; mopup_inv (Suc 0, l, r) lm n ires; wf abc_mopup_LE\ + \ \na. \ (\(s, l, r) n. s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na) n \ + ((t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) + (Suc na), n), + t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) + na, n) \ abc_mopup_LE" +apply(rule allI, rule impI, simp add: t_steps_ind) +apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires") +apply(case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na)", simp) +proof - + fix na a b c + assume "n < length lm" "mopup_inv (a, b, c) lm n ires" "0 < a" + thus "((t_step (a, b, c) (mop_bef n @ tshift mp_up (2 * n), 0), n), + (a, b, c), n) \ abc_mopup_LE" + apply(auto split:if_splits simp add:t_step.simps mopup_inv.simps, + tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *}) + apply(simp_all add: mopupfetchs new_tape.simps abc_mopup_LE_def + lex_triple_def lex_pair_def ) + done +next + fix na + assume "n < length lm" "mopup_inv (Suc 0, l, r) lm n ires" + thus "mopup_inv (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires" + apply(rule mopup_inv_steps) + done +qed + +lemma mopup_halt: "\n < length lm; crsp_l ly (as, lm) (s, l, r) ires\ \ + \ stp. (\ (s, l, r). s = 0) (t_steps (Suc 0, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)" +apply(subgoal_tac "mopup_inv (Suc 0, l, r) lm n ires") +apply(insert wf_abc_mopup_le) +apply(insert halt_lemma[of abc_mopup_LE + "\ ((s, l, r), n). s = 0" + "\ stp. (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)) + , 0) stp, n)"], auto) +apply(insert mopup_halt_pre[of n lm l r], simp, erule exE) +apply(rule_tac x = na in exI, case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) na)", simp) +apply(rule_tac mopup_init, auto) +done +(***End: mopup stop****) + +lemma mopup_halt_conf_pre: + "\n < length lm; crsp_l ly (as, lm) (s, l, r) ires\ + \ \ na. (\ (s', l', r'). s' = 0 \ mopup_stop (s', l', r') lm n ires) + (t_steps (Suc 0, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) na)" +apply(subgoal_tac "\ stp. (\ (s, l, r). s = 0) + (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)", + erule exE) +apply(rule_tac x = stp in exI, + case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp) +apply(subgoal_tac " mopup_inv (Suc 0, l, r) lm n ires") +apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp) lm n ires", simp) +apply(simp only: mopup_inv.simps) +apply(rule_tac mopup_inv_steps, simp, simp) +apply(rule_tac mopup_init, simp, simp) +apply(rule_tac mopup_halt, simp, simp) +done + +lemma mopup_halt_conf: + assumes len: "n < length lm" + and correspond: "crsp_l ly (as, lm) (s, l, r) ires" + shows + "\ na. (\ (s', l', r'). ((\ln rn. s' = 0 \ l' = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires + \ r' = Oc\<^bsup>Suc (abc_lm_v lm n)\<^esup> @ Bk\<^bsup>rn\<^esup>))) + (t_steps (Suc 0, l, r) + ((mop_bef n @ tshift mp_up (2 * n)), 0) na)" +using len correspond mopup_halt_conf_pre[of n lm ly as s l r ires] +apply(simp add: mopup_stop.simps tape_of_nat_abv tape_of_nat_list.simps) +done + +subsection {* Final results about Abacus machine *} + +lemma mopup_halt_bef: "\n < length lm; crsp_l ly (as, lm) (s, l, r) ires\ + \ \stp. (\(s, l, r). s \ 0 \ ((\ (s', l', r'). s' = 0) + (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)))) + (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)" +apply(insert mopup_halt[of n lm ly as s l r ires], simp, erule_tac exE) +proof - + fix stp + assume "n < length lm" + "crsp_l ly (as, lm) (s, l, r) ires" + "(\(s, l, r). s = 0) + (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + thus "\stp. (\(s, ab). 0 < s \ (\(s', l', r'). s' = 0) + (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) + (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + proof(induct stp, simp add: t_steps.simps, simp) + fix stpa + assume h1: + "(\(s, l, r). s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa) \ + \stp. (\(s, ab). 0 < s \ (\(s', l', r'). s' = 0) + (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) + (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + and h2: + "(\(s, l, r). s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) (Suc stpa))" + "n < length lm" + "crsp_l ly (as, lm) (s, l, r) ires" + thus "\stp. (\(s, ab). 0 < s \ (\(s', l', r'). s' = 0) + (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) ( + t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + apply(case_tac "(\(s, l, r). s = 0) (t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", + simp) + apply(rule_tac x = "stpa" in exI) + apply(case_tac "(t_steps (Suc 0, l, r) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", + simp add: t_steps_ind) + done + qed +qed + +lemma tshift_nth_state0: "\n < length tp; tp ! n = (a, 0)\ + \ tshift tp off ! n = (a, 0)" +apply(induct n, case_tac tp, simp) +apply(auto simp: tshift.simps) +done + +lemma shift_length: "length (tshift tp n) = length tp" +apply(auto simp: tshift.simps) +done + +lemma even_Suc_le: "\y mod 2 = 0; 2 * x < y\ \ Suc (2 * x) < y" +by arith + +lemma [simp]: "(4::nat) * n mod 2 = 0" +by arith + +lemma tm_append_fetch_equal: + "\t_ncorrect (tm_of aprog); s'> 0; + fetch (mop_bef n @ tshift mp_up (2 * n)) s' b = (a, 0)\ +\ fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) b + = (a, 0)" +apply(case_tac s', simp) +apply(auto simp: fetch.simps nth_of.simps t_ncorrect.simps shift_length nth_append + tshift.simps split: list.splits block.splits split: if_splits) +done + +lemma [simp]: + "\t_ncorrect (tm_of aprog); + t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) = + (0, l'', r''); s' > 0\ + \ t_step (s' + length (tm_of aprog) div 2, l', r') + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2), 0) = (0, l'', r'')" +apply(simp add: t_step.simps) +apply(subgoal_tac + "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' + (case r' of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc)) + = (fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) + (case r' of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", simp) +apply(case_tac "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' + (case r' of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc))", simp) +apply(drule_tac tm_append_fetch_equal, auto) +done + +lemma [intro]: + "start_of (layout_of aprog) (length aprog) - Suc 0 = + length (tm_of aprog) div 2" +apply(subgoal_tac "abc2t_correct aprog") +apply(insert pre_lheq[of "concat (take (length aprog) + (tms_of aprog))" "length aprog" aprog], simp add: tm_of.simps) +by auto + +lemma tm_append_stop_step: + "\t_ncorrect (tm_of aprog); + t_ncorrect (mop_bef n @ tshift mp_up (2 * n)); n < length lm; + (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp) = + (s', l', r'); + s' \ 0; + t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) + = (0, l'', r'')\ + \ +(t_steps ((start_of (layout_of aprog) (length aprog), l, r)) + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) (Suc stp)) + = (0, l'', r'')" +apply(insert tm_append_steps_equal[of "tm_of aprog" + "(mop_bef n @ tshift mp_up (2 * n))" + "(start_of (layout_of aprog) (length aprog) - Suc 0)" + "Suc 0" l r stp], simp) +apply(subgoal_tac "(start_of (layout_of aprog) (length aprog) - Suc 0) + = (length (tm_of aprog) div 2)", simp add: t_steps_ind) +apply(case_tac + "(t_steps (start_of (layout_of aprog) (length aprog), l, r) + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (length (tm_of aprog) div 2), 0) stp)", simp) +apply(subgoal_tac "start_of (layout_of aprog) (length aprog) > 0", + case_tac "start_of (layout_of aprog) (length aprog)", + simp, simp) +apply(rule startof_not0, auto) +done + +lemma start_of_out_range: +"as \ length aprog \ + start_of (layout_of aprog) as = + start_of (layout_of aprog) (length aprog)" +apply(induct as, simp) +apply(case_tac "length aprog = Suc as", simp) +apply(simp add: start_of.simps) +done + +lemma [intro]: "t_ncorrect (tm_of aprog)" +apply(simp add: tm_of.simps) +apply(insert tms_mod2[of "length aprog" aprog], + simp add: t_ncorrect.simps) +done + +lemma abacus_turing_eq_halt_case_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly ac tc ires; + n < length am; + abc_steps_l ac aprog stp = (as, am); + mop_ss = start_of ly (length aprog); + as \ length aprog\ + \ \ stp. (\ (s, l, r). s = 0 \ mopup_inv (0, l, r) am n ires) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)" +apply(insert steps_crsp[of ly aprog tprog ac tc ires "(as, am)" stp], auto) +apply(case_tac "(t_steps tc (tm_of aprog, 0) n')", + simp add: tMp.simps) +apply(subgoal_tac "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))") +proof - + fix n' a b c + assume h1: + "crsp_l (layout_of aprog) ac tc ires" + "abc_steps_l ac aprog stp = (as, am)" + "length aprog \ as" + "crsp_l (layout_of aprog) (as, am) (a, b, c) ires" + "t_steps tc (tm_of aprog, 0) n' = (a, b, c)" + "n < length am" + "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))" + hence h2: + "t_steps tc (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) n' + = (a, b, c)" + apply(rule_tac tm_append_steps, simp) + apply(simp add: crsp_l.simps, auto) + apply(simp add: crsp_l.simps) + apply(rule startof_not0) + done + from h1 have h3: + "\stp. (\(s, l, r). s \ 0 \ ((\ (s', l', r'). s' = 0) + (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)))) + (t_steps (Suc 0, b, c) + (mop_bef n @ tshift mp_up (2 * n), 0) stp)" + apply(rule_tac mopup_halt_bef, auto) + done + from h1 and h2 and h3 show + "\stp. case t_steps tc (tm_of aprog @ abacus.tshift (mop_bef n @ abacus.tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) stp of (s, ab) + \ s = 0 \ mopup_inv (0, ab) am n ires" + proof(erule_tac exE, + case_tac "(t_steps (Suc 0, b, c) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", simp, + case_tac "(t_step (aa, ba, ca) + (mop_bef n @ tshift mp_up (2 * n), 0))", simp) + fix stpa aa ba ca aaa baa caa + assume g1: "0 < aa \ aaa = 0" + "t_steps (Suc 0, b, c) + (mop_bef n @ tshift mp_up (2 * n), 0) stpa = (aa, ba,ca)" + "t_step (aa, ba, ca) (mop_bef n @ tshift mp_up (2 * n), 0) + = (0, baa, caa)" + from h1 and this have g2: + "t_steps (start_of (layout_of aprog) (length aprog), b, c) + (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) + (Suc stpa) = (0, baa, caa)" + apply(rule_tac tm_append_stop_step, auto) + done + from h1 and h2 and g1 and this show "?thesis" + apply(rule_tac x = "n' + Suc stpa" in exI) + apply(simp add: t_steps_ind del: t_steps.simps) + apply(subgoal_tac "a = start_of (layout_of aprog) + (length aprog)", simp) + apply(insert mopup_inv_steps[of n am "Suc 0" b c ires "Suc stpa"], + simp add: t_steps_ind) + apply(subgoal_tac "mopup_inv (Suc 0, b, c) am n ires", simp) + apply(rule_tac mopup_init, simp, simp) + apply(simp add: crsp_l.simps) + apply(erule_tac start_of_out_range) + done + qed +next + show " t_ncorrect (mop_bef n @ tshift mp_up (2 * n))" + apply(auto simp: t_ncorrect.simps tshift.simps mp_up_def) + done +qed + +text {* + One of the main theorems about Abacus compilation. +*} + +lemma abacus_turing_eq_halt_case: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + "ly = layout_of aprog" + and complied: + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + "tprog = tm_of aprog" + and correspond: + -- {* TM configuration @{text "tc"} and Abacus configuration @{text "ac"} + are in correspondence: *} + "crsp_l ly ac tc ires" + and halt_state: + -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So + if Abacus is in such a state, it is in halt state: *} + "as \ length aprog" + and abc_exec: + -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"} + reaches such a halt state: *} + "abc_steps_l ac aprog stp = (as, am)" + and rs_len: + -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *} + "n < length am" + and mopup_start: + -- {* The startling label for mopup mahines, according to the layout and Abacus program + should be @{text "mop_ss"}: *} + "mop_ss = start_of ly (length aprog)" + shows + -- {* + After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup + TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which + only hold the content of memory cell @{text "n"}: + *} + "\ stp. (\ (s, l, r). \ ln rn. s = 0 \ l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires + \ r = Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>rn\<^esup>) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)" +proof - + from layout complied correspond halt_state abc_exec rs_len mopup_start + and abacus_turing_eq_halt_case_pre [of ly aprog tprog ac tc ires n am stp as mop_ss] + show "?thesis" + apply(simp add: mopup_inv.simps mopup_stop.simps tape_of_nat_abv) + done +qed + +lemma abc_unhalt_case_zero: +"\crsp_l (layout_of aprog) ac tc ires; + n < length am; + \stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp)\ + \ \astp bstp. 0 \ bstp \ + crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires" +apply(rule_tac x = "Suc 0" in exI) +apply(case_tac " abc_steps_l ac aprog (Suc 0)", simp) +proof - + fix a b + assume "crsp_l (layout_of aprog) ac tc ires" + "abc_steps_l ac aprog (Suc 0) = (a, b)" + thus "\bstp. crsp_l (layout_of aprog) (a, b) + (t_steps tc (tm_of aprog, 0) bstp) ires" + apply(insert steps_crsp[of "layout_of aprog" aprog + "tm_of aprog" ac tc ires "(a, b)" "Suc 0"], auto) + done +qed + +declare abc_steps_l.simps[simp del] + +lemma abc_steps_ind: + "let (as, am) = abc_steps_l ac aprog stp in + abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog) " +proof(simp) + show "(\(as, am). abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog stp)" + proof(induct stp arbitrary: ac) + fix ac + show "(\(as, am). abc_steps_l ac aprog (Suc 0) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog 0)" + apply(case_tac "ac:: nat \ nat list", + simp add: abc_steps_l.simps) + apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", + simp add: abc_steps_l.simps) + done + next + fix stp ac + assume h1: + "(\ac. (\(as, am). abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog stp))" + thus + "(\(as, am). abc_steps_l ac aprog (Suc (Suc stp)) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l ac aprog (Suc stp))" + apply(case_tac "ac::nat \ nat list", simp) + apply(subgoal_tac + "abc_steps_l (a, b) aprog (Suc (Suc stp)) = + abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) + aprog (Suc stp)", simp) + apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", simp) + proof - + fix a b aa ba + assume h2: "abc_step_l (a, b) (abc_fetch a aprog) = (aa, ba)" + from h1 and h2 show + "(\(as, am). abc_steps_l (aa, ba) aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)) + (abc_steps_l (a, b) aprog (Suc stp))" + apply(insert h1[of "(aa, ba)"]) + apply(simp add: abc_steps_l.simps) + apply(insert h2, simp) + done + next + fix a b + show + "abc_steps_l (a, b) aprog (Suc (Suc stp)) = + abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) + aprog (Suc stp)" + apply(simp only: abc_steps_l.simps) + done + qed + qed +qed + +lemma abc_unhalt_case_induct: + "\crsp_l (layout_of aprog) ac tc ires; + n < length am; + \stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp); + stp \ bstp; + crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires\ + \ \astp bstp. Suc stp \ bstp \ crsp_l (layout_of aprog) + (abc_steps_l ac aprog astp) (t_steps tc (tm_of aprog, 0) bstp) ires" +apply(rule_tac x = "Suc astp" in exI) +apply(case_tac "abc_steps_l ac aprog astp") +proof - + fix a b + assume + "\stp. (\(as, am). as < length aprog) + (abc_steps_l ac aprog stp)" + "stp \ bstp" + "crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires" + "abc_steps_l ac aprog astp = (a, b)" + thus + "\bstp\Suc stp. crsp_l (layout_of aprog) + (abc_steps_l ac aprog (Suc astp)) + (t_steps tc (tm_of aprog, 0) bstp) ires" + apply(insert crsp_inside[of "layout_of aprog" aprog + "tm_of aprog" a b "(t_steps tc (tm_of aprog, 0) bstp)" "ires"], auto) + apply(erule_tac x = astp in allE, auto) + apply(rule_tac x = "bstp + stpa" in exI, simp) + apply(insert abc_steps_ind[of ac aprog "astp"], simp) + done +qed + +lemma abc_unhalt_case: + "\crsp_l (layout_of aprog) ac tc ires; + \stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp)\ + \ (\ astp bstp. bstp \ stp \ + crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) + (t_steps tc (tm_of aprog, 0) bstp) ires)" +apply(induct stp) +apply(rule_tac abc_unhalt_case_zero, auto) +apply(rule_tac abc_unhalt_case_induct, auto) +done + +lemma abacus_turing_eq_unhalt_case_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + crsp_l ly ac tc ires; + \ stp. ((\ (as, am). as < length aprog) + (abc_steps_l ac aprog stp)); + mop_ss = start_of ly (length aprog)\ + \ (\ (\ stp. (\ (s, l, r). s = 0) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)))" + apply(auto) +proof - + fix stp a b + assume h1: + "crsp_l (layout_of aprog) ac tc ires" + "\stp. (\(as, am). as < length aprog) (abc_steps_l ac aprog stp)" + "t_steps tc (tm_of aprog @ tMp n (start_of (layout_of aprog) + (length aprog) - Suc 0), 0) stp = (0, a, b)" + thus "False" + proof(insert abc_unhalt_case[of aprog ac tc ires stp], auto, + case_tac "(abc_steps_l ac aprog astp)", + case_tac "(t_steps tc (tm_of aprog, 0) bstp)", simp) + fix astp bstp aa ba aaa baa c + assume h2: + "abc_steps_l ac aprog astp = (aa, ba)" "stp \ bstp" + "t_steps tc (tm_of aprog, 0) bstp = (aaa, baa, c)" + "crsp_l (layout_of aprog) (aa, ba) (aaa, baa, c) ires" + hence h3: + "t_steps tc (tm_of aprog @ tMp n + (start_of (layout_of aprog) (length aprog) - Suc 0), 0) bstp + = (aaa, baa, c)" + apply(intro tm_append_steps, auto) + apply(simp add: crsp_l.simps, rule startof_not0) + done + from h2 have h4: "\ diff. bstp = stp + diff" + apply(rule_tac x = "bstp - stp" in exI, simp) + done + from h4 and h3 and h2 and h1 show "?thesis" + apply(auto) + apply(simp add: state0_ind crsp_l.simps) + apply(subgoal_tac "start_of (layout_of aprog) aa > 0", simp) + apply(rule startof_not0) + done + qed +qed + +lemma abacus_turing_eq_unhalt_case: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + "ly = layout_of aprog" + and compiled: + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + "tprog = tm_of aprog" + and correspond: + -- {* + TM configuration @{text "tc"} and Abacus configuration @{text "ac"} + are in correspondence: + *} + "crsp_l ly ac tc ires" + and abc_unhalt: + -- {* + If, no matter how many steps the Abacus program @{text "aprog"} executes, it + may never reach a halt state. + *} + "\ stp. ((\ (as, am). as < length aprog) + (abc_steps_l ac aprog stp))" + and mopup_start: "mop_ss = start_of ly (length aprog)" + shows + -- {* + The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well. + *} + "\ (\ stp. (\ (s, l, r). s = 0) + (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp))" + using layout compiled correspond abc_unhalt mopup_start + apply(rule_tac abacus_turing_eq_unhalt_case_pre, auto) + done + + +definition abc_list_crsp:: "nat list \ nat list \ bool" + where + "abc_list_crsp xs ys = (\ n. xs = ys @ 0\<^bsup>n\<^esup> \ ys = xs @ 0\<^bsup>n\<^esup>)" +lemma [intro]: "abc_list_crsp (lm @ 0\<^bsup>m\<^esup>) lm" +apply(auto simp: abc_list_crsp_def) +done + +lemma abc_list_crsp_lm_v: + "abc_list_crsp lma lmb \ abc_lm_v lma n = abc_lm_v lmb n" +apply(auto simp: abc_list_crsp_def abc_lm_v.simps + nth_append exponent_def) +done + +lemma rep_app_cons_iff: + "k < n \ replicate n a[k:=b] = + replicate k a @ b # replicate (n - k - 1) a" +apply(induct n arbitrary: k, simp) +apply(simp split:nat.splits) +done + +lemma abc_list_crsp_lm_s: + "abc_list_crsp lma lmb \ + abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)" +apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps) +apply(simp_all add: list_update_append, auto simp: exponent_def) +proof - + fix na + assume h: "m < length lmb + na" " \ m < length lmb" + hence "m - length lmb < na" by simp + hence "replicate na 0[(m- length lmb):= n] = + replicate (m - length lmb) 0 @ n # + replicate (na - (m - length lmb) - 1) 0" + apply(erule_tac rep_app_cons_iff) + done + thus "\nb. replicate na 0[m - length lmb := n] = + replicate (m - length lmb) 0 @ n # replicate nb 0 \ + replicate (m - length lmb) 0 @ [n] = + replicate na 0[m - length lmb := n] @ replicate nb 0" + apply(auto) + done +next + fix na + assume h: "\ m < length lmb + na" + show + "\nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] = + replicate (m - length lmb) 0 @ n # replicate nb 0 \ + replicate (m - length lmb) 0 @ [n] = + replicate na 0 @ + replicate (m - (length lmb + na)) 0 @ n # replicate nb 0" + apply(rule_tac x = 0 in exI, simp, auto) + using h + apply(simp add: replicate_add[THEN sym]) + done +next + fix na + assume h: "\ m < length lma" "m < length lma + na" + hence "m - length lma < na" by simp + hence + "replicate na 0[(m- length lma):= n] = replicate (m - length lma) + 0 @ n # replicate (na - (m - length lma) - 1) 0" + apply(erule_tac rep_app_cons_iff) + done + thus "\nb. replicate (m - length lma) 0 @ [n] = + replicate na 0[m - length lma := n] @ replicate nb 0 + \ replicate na 0[m - length lma := n] = + replicate (m - length lma) 0 @ n # replicate nb 0" + apply(auto) + done +next + fix na + assume "\ m < length lma + na" + thus " \nb. replicate (m - length lma) 0 @ [n] = + replicate na 0 @ + replicate (m - (length lma + na)) 0 @ n # replicate nb 0 + \ replicate na 0 @ + replicate (m - (length lma + na)) 0 @ [n] = + replicate (m - length lma) 0 @ n # replicate nb 0" + apply(rule_tac x = 0 in exI, simp, auto) + apply(simp add: replicate_add[THEN sym]) + done +qed + +lemma abc_list_crsp_step: + "\abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); + abc_step_l (aa, lmb) i = (a', lmb')\ + \ a' = a \ abc_list_crsp lma' lmb'" +apply(case_tac i, auto simp: abc_step_l.simps + abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def + split: abc_inst.splits if_splits) +done + +lemma abc_steps_red: + "abc_steps_l ac aprog stp = (as, am) \ + abc_steps_l ac aprog (Suc stp) = + abc_step_l (as, am) (abc_fetch as aprog)" +using abc_steps_ind[of ac aprog stp] +apply(simp) +done + +lemma abc_list_crsp_steps: + "\abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (a, lm'); aprog \ []\ + \ \ lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ + abc_list_crsp lm' lma" +apply(induct stp arbitrary: a lm', simp add: abc_steps_l.simps, auto) +apply(case_tac "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp", + simp add: abc_steps_ind) +proof - + fix stp a lm' aa b + assume ind: + "\a lm'. aa = a \ b = lm' \ + \lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ + abc_list_crsp lm' lma" + and h: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) = (a, lm')" + "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (aa, b)" + "aprog \ []" + hence g1: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) + = abc_step_l (aa, b) (abc_fetch aa aprog)" + apply(rule_tac abc_steps_red, simp) + done + have "\lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \ + abc_list_crsp b lma" + apply(rule_tac ind, simp) + done + from this obtain lma where g2: + "abc_steps_l (0, lm) aprog stp = (aa, lma) \ + abc_list_crsp b lma" .. + hence g3: "abc_steps_l (0, lm) aprog (Suc stp) + = abc_step_l (aa, lma) (abc_fetch aa aprog)" + apply(rule_tac abc_steps_red, simp) + done + show "\lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \ + abc_list_crsp lm' lma" + using g1 g2 g3 h + apply(auto) + apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)", + case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp) + apply(rule_tac abc_list_crsp_step, auto) + done +qed + +lemma [simp]: "(case ca of [] \ Bk | Bk # xs \ Bk | Oc # xs \ Oc) = + (case ca of [] \ Bk | x # xs \ x)" +by(case_tac ca, simp_all, case_tac a, simp, simp) + +lemma steps_eq: "length t mod 2 = 0 \ + t_steps c (t, 0) stp = steps c t stp" +apply(induct stp) +apply(simp add: steps.simps t_steps.simps) +apply(simp add:tstep_red t_steps_ind) +apply(case_tac "steps c t stp", simp) +apply(auto simp: t_step.simps tstep.simps) +done + +lemma crsp_l_start: "crsp_l ly (0, lm) (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) ires" +apply(simp add: crsp_l.simps, auto simp: start_of.simps) +done + +lemma t_ncorrect_app: "\t_ncorrect t1; t_ncorrect t2\ \ + t_ncorrect (t1 @ t2)" +apply(simp add: t_ncorrect.simps, auto) +done + +lemma [simp]: + "(length (tm_of aprog) + + length (tMp n (start_of ly (length aprog) - Suc 0))) mod 2 = 0" +apply(subgoal_tac + "t_ncorrect (tm_of aprog @ tMp n + (start_of ly (length aprog) - Suc 0))") +apply(simp add: t_ncorrect.simps) +apply(rule_tac t_ncorrect_app, + auto simp: tMp.simps t_ncorrect.simps tshift.simps mp_up_def) +apply(subgoal_tac + "t_ncorrect (tm_of aprog)", simp add: t_ncorrect.simps) +apply(auto) +done + +lemma [simp]: "takeWhile (\a. a = Oc) + (replicate rs Oc @ replicate rn Bk) = replicate rs Oc" +apply(induct rs, auto) +apply(induct rn, auto) +done + +lemma abacus_turing_eq_halt': + "\ly = layout_of aprog; + tprog = tm_of aprog; + n < length am; + abc_steps_l (0, lm) aprog stp = (as, am); + mop_ss = start_of ly (length aprog); + as \ length aprog\ + \ \ stp m l. steps (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) + (tprog @ (tMp n (mop_ss - 1))) stp + = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" +apply(drule_tac tc = "(Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>)" in + abacus_turing_eq_halt_case, auto intro: crsp_l_start) +apply(subgoal_tac + "length (tm_of aprog @ tMp n + (start_of ly (length aprog) - Suc 0)) mod 2 = 0") +apply(simp add: steps_eq) +apply(rule_tac x = stpa in exI, + simp add: exponent_def, auto) +done + + +lemma list_length: "xs = ys \ length xs = length ys" +by simp +lemma [elim]: "tinres (Bk\<^bsup>m\<^esup>) b \ \m. b = Bk\<^bsup>m\<^esup>" +apply(auto simp: tinres_def) +apply(rule_tac x = "m-n" in exI, + auto simp: exponent_def replicate_add[THEN sym]) +apply(case_tac "m < n", auto) +apply(drule_tac list_length, auto) +apply(subgoal_tac "\ d. m = d + n", auto simp: replicate_add) +apply(rule_tac x = "m - n" in exI, simp) +done +lemma [intro]: "tinres [Bk] (Bk\<^bsup>k\<^esup>) " +apply(auto simp: tinres_def exponent_def) +apply(case_tac k, auto) +apply(rule_tac x = "Suc 0" in exI, simp) +done + +lemma abacus_turing_eq_halt_pre: + "\ly = layout_of aprog; + tprog = tm_of aprog; + n < length am; + abc_steps_l (0, lm) aprog stp = (as, am); + mop_ss = start_of ly (length aprog); + as \ length aprog\ + \ \ stp m l. steps (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) + (tprog @ (tMp n (mop_ss - 1))) stp + = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" +using abacus_turing_eq_halt' +apply(simp) +done + + +text {* + Main theorem for the case when the original Abacus program does halt. +*} + +lemma abacus_turing_eq_halt: + assumes layout: + "ly = layout_of aprog" + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + and compiled: "tprog = tm_of aprog" + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + and halt_state: + -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So + if Abacus is in such a state, it is in halt state: *} + "as \ length aprog" + and abc_exec: + -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"} + reaches such a halt state: *} + "abc_steps_l (0, lm) aprog stp = (as, am)" + and rs_locate: + -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *} + "n < length am" + and mopup_start: + -- {* The startling label for mopup mahines, according to the layout and Abacus program + should be @{text "mop_ss"}: *} + "mop_ss = start_of ly (length aprog)" + shows + -- {* + After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup + TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which + only hold the content of memory cell @{text "n"}: + *} + "\ stp m l. steps (Suc 0, Bk # Bk # ires, @ Bk\<^bsup>rn\<^esup>) (tprog @ (tMp n (mop_ss - 1))) stp + = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" + using layout compiled halt_state abc_exec rs_locate mopup_start + by(rule_tac abacus_turing_eq_halt_pre, auto) + +lemma abacus_turing_eq_uhalt': + "\ly = layout_of aprog; + tprog = tm_of aprog; + \ stp. ((\ (as, am). as < length aprog) + (abc_steps_l (0, lm) aprog stp)); + mop_ss = start_of ly (length aprog)\ + \ (\ (\ stp. isS0 (steps (Suc 0, [Bk, Bk], ) + (tprog @ (tMp n (mop_ss - 1))) stp)))" +apply(drule_tac tc = "(Suc 0, [Bk, Bk], )" and n = n and ires = "[]" in + abacus_turing_eq_unhalt_case, auto intro: crsp_l_start) +apply(simp add: crsp_l.simps start_of.simps) +apply(erule_tac x = stp in allE, erule_tac x = stp in allE) +apply(subgoal_tac + "length (tm_of aprog @ tMp n + (start_of ly (length aprog) - Suc 0)) mod 2 = 0") +apply(simp add: steps_eq, auto simp: isS0_def) +done + +text {* + Main theorem for the case when the original Abacus program does not halt. + *} +lemma abacus_turing_eq_uhalt: + assumes layout: + -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} + "ly = layout_of aprog" + and compiled: + -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} + "tprog = tm_of aprog" + and abc_unhalt: + -- {* + If, no matter how many steps the Abacus program @{text "aprog"} executes, it + may never reach a halt state. + *} + "\ stp. ((\ (as, am). as < length aprog) + (abc_steps_l (0, lm) aprog stp))" + and mop_start: "mop_ss = start_of ly (length aprog)" + shows + -- {* + The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well. + *} + "\ (\ stp. isS0 (steps (Suc 0, [Bk, Bk], ) + (tprog @ (tMp n (mop_ss - 1))) stp))" + using abacus_turing_eq_uhalt' + layout compiled abc_unhalt mop_start + by(auto) + +end +