diff --git a/thys/Cook_Levin/Arithmetic.thy b/thys/Cook_Levin/Arithmetic.thy --- a/thys/Cook_Levin/Arithmetic.thy +++ b/thys/Cook_Levin/Arithmetic.thy @@ -1,6500 +1,6497 @@ section \Arithmetic\label{s:tm-arithmetic}\ theory Arithmetic imports Memorizing begin text \ In this section we define a representation of natural numbers and some reusable Turing machines for elementary arithmetical operations. All Turing machines implementing the operations assume that the tape heads on the tapes containing the operands and the result(s) contain one natural number each. In programming language terms we could say that such a tape corresponds to a variable of type @{typ nat}. Furthermore, initially the tape heads are on cell number~1, that is, one to the right of the start symbol. The Turing machines will halt with the tape heads in that position as well. In that way operations can be concatenated seamlessly. \ subsection \Binary numbers\label{s:tm-arithmetic-binary}\ text \ We represent binary numbers as sequences of the symbols \textbf{0} and \textbf{1}. Slightly unusually the least significant bit will be on the left. While every sequence over these symbols represents a natural number, the representation is not unique due to leading (or rather, trailing) zeros. The \emph{canonical} representation is unique and has no trailing zeros, not even for the number zero, which is thus represented by the empty symbol sequence. As a side effect empty tapes can be thought of as being initialized with zero. Naturally the binary digits 0 and 1 are represented by the symbols \textbf{0} and \textbf{1}, respectively. For example, the decimal number $14$, conventionally written $1100_2$ in binary, is represented by the symbol sequence \textbf{0011}. The next two functions map between symbols and binary digits: \ abbreviation (input) tosym :: "nat \ symbol" where "tosym z \ z + 2" abbreviation todigit :: "symbol \ nat" where "todigit z \ if z = \ then 1 else 0" text \ The numerical value of a symbol sequence: \ definition num :: "symbol list \ nat" where "num xs \ \i\[0.. The $i$-th digit of a symbol sequence, where digits out of bounds are considered trailing zeros: \ definition digit :: "symbol list \ nat \ nat" where "digit xs i \ if i < length xs then xs ! i else 0" text \ Some properties of $num$: \ lemma num_ge_pow: assumes "i < length xs" and "xs ! i = \" shows "num xs \ 2 ^ i" proof - let ?ys = "map (\i. todigit (xs ! i) * 2 ^ i) [0.. 2 ^ i" unfolding num_def using elem_le_sum_list by (metis (no_types, lifting)) qed lemma num_trailing_zero: assumes "todigit z = 0" shows "num xs = num (xs @ [z])" proof - let ?xs = "xs @ [z]" let ?ys = "map (\i. todigit (?xs ! i) * 2 ^ i) [0..i. todigit (xs ! i) * 2 ^ i) [0..i. todigit (xs ! i) * 2 ^ i) [0..i. todigit (xs ! i) * 2 ^ i) [0..i. todigit ((x # xs) ! i) * 2 ^ i) [0..i. todigit ((x # xs) ! i) * 2 ^ i) [0..<1]) @ (map (\i. todigit ((x # xs) ! i) * 2 ^ i) [1..i. f i) [1..i. f (Suc i)) [0.. nat" and m proof (rule nth_equalityI) show "length (map f [1..i. f (Suc i)) [0..i. i < length (map f [1.. map f [1..i. f (Suc i)) [0..i\[1..i\[0.. nat" and m by simp have "num (x # xs) = (\i\[0..i\[0..<1]. (todigit ((x # xs) ! i) * 2 ^ i)) + (\i\[1..i\[1..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0.. A symbol sequence is a canonical representation of a natural number if the sequence contains only the symbols \textbf{0} and \textbf{1} and is either empty or ends in \textbf{1}. \ definition canonical :: "symbol list \ bool" where "canonical xs \ bit_symbols xs \ (xs = [] \ last xs = \)" lemma canonical_Cons: assumes "canonical xs" and "xs \ []" and "x = \ \ x = \" shows "canonical (x # xs)" using assms canonical_def less_Suc_eq_0_disj by auto lemma canonical_Cons_3: "canonical xs \ canonical (\ # xs)" using canonical_def less_Suc_eq_0_disj by auto lemma canonical_tl: "canonical (x # xs) \ canonical xs" using canonical_def by fastforce lemma prepend_2_even: "x = \ \ even (num (x # xs))" using num_Cons by simp lemma prepend_3_odd: "x = \ \ odd (num (x # xs))" using num_Cons by simp text \ Every number has exactly one canonical representation. \ lemma canonical_ex1: fixes n :: nat shows "\!xs. num xs = n \ canonical xs" proof (induction n rule: nat_less_induct) case IH: (1 n) show ?case proof (cases "n = 0") case True have "num [] = 0" using num_def by simp moreover have "canonical xs \ num xs = 0 \ xs = []" for xs proof (rule ccontr) fix xs assume "canonical xs" "num xs = 0" "xs \ []" then have "length xs > 0" "last xs = \" using canonical_def by simp_all then have "xs ! (length xs - 1) = \" by (metis Suc_diff_1 last_length) then have "num xs \ 2 ^ (length xs - 1)" using num_ge_pow by (meson \0 < length xs\ diff_less zero_less_one) then have "num xs > 0" by (meson dual_order.strict_trans1 le0 le_less_trans less_exp) then show False using \num xs = 0\ by auto qed ultimately show ?thesis using IH True canonical_def by (metis less_nat_zero_code list.size(3)) next case False then have gt: "n > 0" by simp define m where "m = n div 2" define r where "r = n mod 2" have n: "n = 2 * m + r" using m_def r_def by simp have "m < n" using gt m_def by simp then obtain xs where "num xs = m" "canonical xs" using IH by auto then have "num (tosym r # xs) = n" (is "num ?xs = n") using num_Cons n add.commute r_def by simp have "canonical ?xs" proof (cases "r = 0") case True then have "m > 0" using gt n by simp then have "xs \ []" using `num xs = m` num_def by auto then show ?thesis using canonical_Cons[of xs] `canonical xs` r_def True by simp next case False then show ?thesis using `canonical xs` canonical_Cons_3 r_def by (metis One_nat_def not_mod_2_eq_1_eq_0 numeral_3_eq_3 one_add_one plus_1_eq_Suc) qed moreover have "xs1 = xs2" if "canonical xs1" "num xs1 = n" "canonical xs2" "num xs2 = n" for xs1 xs2 proof - have "xs1 \ []" using gt that(2) num_def by auto then obtain x1 ys1 where 1: "xs1 = x1 # ys1" by (meson neq_Nil_conv) then have x1: "x1 = \ \ x1 = \" using canonical_def that(1) by auto have "xs2 \ []" using gt that(4) num_def by auto then obtain x2 ys2 where 2: "xs2 = x2 # ys2" by (meson neq_Nil_conv) then have x2: "x2 = \ \ x2 = \" using canonical_def that(3) by auto have "x1 = x2" using prepend_2_even prepend_3_odd that 1 2 x1 x2 by metis moreover have "n = todigit x1 + 2 * num ys1" using that(2) num_Cons 1 by simp moreover have "n = todigit x2 + 2 * num ys2" using that(4) num_Cons 2 by simp ultimately have "num ys1 = num ys2" by simp moreover have "num ys1 < n" using that(2) num_Cons 1 gt by simp moreover have "num ys2 < n" using that(4) num_Cons 2 gt by simp ultimately have "ys1 = ys2" using IH 1 2 that(1,3) by (metis canonical_tl) then show "xs1 = xs2" using `x1 = x2` 1 2 by simp qed ultimately show ?thesis using \num (tosym r # xs) = n\ by auto qed qed text \ The canonical representation of a natural number as symbol sequence: \ definition canrepr :: "nat \ symbol list" where "canrepr n \ THE xs. num xs = n \ canonical xs" lemma canrepr_inj: "inj canrepr" using canrepr_def canonical_ex1 by (smt (verit, del_insts) inj_def the_equality) lemma canonical_canrepr: "canonical (canrepr n)" using theI'[OF canonical_ex1] canrepr_def by simp lemma canrepr: "num (canrepr n) = n" using theI'[OF canonical_ex1] canrepr_def by simp lemma bit_symbols_canrepr: "bit_symbols (canrepr n)" using canonical_canrepr canonical_def by simp lemma proper_symbols_canrepr: "proper_symbols (canrepr n)" using bit_symbols_canrepr by fastforce lemma canreprI: "num xs = n \ canonical xs \ canrepr n = xs" using canrepr canonical_canrepr canonical_ex1 by blast lemma canrepr_0: "canrepr 0 = []" using num_def canonical_def by (intro canreprI) simp_all lemma canrepr_1: "canrepr 1 = [\]" using num_def canonical_def by (intro canreprI) simp_all text \ The length of the canonical representation of a number $n$: \ abbreviation nlength :: "nat \ nat" where "nlength n \ length (canrepr n)" lemma nlength_0: "nlength n = 0 \ n = 0" by (metis canrepr canrepr_0 length_0_conv) corollary nlength_0_simp [simp]: "nlength 0 = 0" using nlength_0 by simp lemma num_replicate2_eq_pow: "num (replicate j \ @ [\]) = 2 ^ j" proof (induction j) case 0 then show ?case using num_def by simp next case (Suc j) then show ?case using num_Cons by simp qed lemma num_replicate3_eq_pow_minus_1: "num (replicate j \) = 2 ^ j - 1" proof (induction j) case 0 then show ?case using num_def by simp next case (Suc j) then have "num (replicate (Suc j) \) = num (\ # replicate j \)" by simp also have "... = 1 + 2 * (2 ^ j - 1)" using Suc num_Cons by simp also have "... = 1 + 2 * 2 ^ j - 2" by (metis Nat.add_diff_assoc diff_mult_distrib2 mult_2 mult_le_mono2 nat_1_add_1 one_le_numeral one_le_power) also have "... = 2 ^ Suc j - 1" by simp finally show ?case . qed lemma nlength_pow2: "nlength (2 ^ j) = Suc j" proof - define xs :: "nat list" where "xs = replicate j 2 @ [3]" then have "length xs = Suc j" by simp moreover have "num xs = 2 ^ j" using num_replicate2_eq_pow xs_def by simp moreover have "canonical xs" using xs_def bit_symbols_append canonical_def by simp ultimately show ?thesis using canreprI by blast qed corollary nlength_1_simp [simp]: "nlength 1 = 1" using nlength_pow2[of 0] by simp corollary nlength_2: "nlength 2 = 2" using nlength_pow2[of 1] by simp lemma nlength_pow_minus_1: "nlength (2 ^ j - 1) = j" proof - define xs :: "nat list" where "xs = replicate j \" then have "length xs = j" by simp moreover have "num xs = 2 ^ j - 1" using num_replicate3_eq_pow_minus_1 xs_def by simp moreover have "canonical xs" proof - have "bit_symbols xs" using xs_def by simp moreover have "last xs = 3 \ xs = []" by (cases "j = 0") (simp_all add: xs_def) ultimately show ?thesis using canonical_def by auto qed ultimately show ?thesis using canreprI by metis qed corollary nlength_3: "nlength 3 = 2" using nlength_pow_minus_1[of 2] by simp text \ When handling natural numbers, Turing machines will usually have tape contents of the following form: \ abbreviation ncontents :: "nat \ (nat \ symbol)" ("\_\\<^sub>N") where "\n\\<^sub>N \ \canrepr n\" lemma ncontents_0: "\0\\<^sub>N = \[]\" by (simp add: canrepr_0) lemma clean_tape_ncontents: "clean_tape (\x\\<^sub>N, i)" using bit_symbols_canrepr clean_contents_proper by fastforce lemma ncontents_1_blank_iff_zero: "\n\\<^sub>N 1 = \ \ n = 0" using bit_symbols_canrepr contents_def nlength_0 by (metis contents_outofbounds diff_is_0_eq' leI length_0_conv length_greater_0_conv less_one zero_neq_numeral) text \ Every bit symbol sequence can be turned into a canonical representation of some number by stripping trailing zeros. The length of the prefix without trailing zeros is given by the next function: \ definition canlen :: "symbol list \ nat" where "canlen zs \ LEAST m. \i m \ zs ! i = \" lemma canlen_at_ge: "\i canlen zs \ zs ! i = \" proof - let ?P = "\m. \i m \ zs ! i = \" have "?P (length zs)" by simp then show ?thesis unfolding canlen_def using LeastI[of ?P "length zs"] by fast qed lemma canlen_eqI: assumes "\i m \ zs ! i = \" and "\y. \i y \ zs ! i = \ \ m \ y" shows "canlen zs = m" unfolding canlen_def using assms Least_equality[of _ m, OF _ assms(2)] by presburger lemma canlen_le_length: "canlen zs \ length zs" proof - let ?P = "\m. \i m \ zs ! i = \" have "?P (length zs)" by simp then show ?thesis unfolding canlen_def using Least_le[of _ "length zs"] by simp qed lemma canlen_le: assumes "\i m \ zs ! i = \" shows "m \ canlen zs" unfolding canlen_def using Least_le[of _ m] assms by simp lemma canlen_one: assumes "bit_symbols zs" and "canlen zs > 0" shows "zs ! (canlen zs - 1) = \" proof (rule ccontr) assume "zs ! (canlen zs - 1) \ \" then have "zs ! (canlen zs - 1) = \" using assms canlen_le_length by (metis One_nat_def Suc_pred lessI less_le_trans) then have "\i canlen zs - 1 \ zs ! i = 2" using canlen_at_ge assms(2) by (metis One_nat_def Suc_leI Suc_pred le_eq_less_or_eq) then have "canlen zs - 1 \ canlen zs" using canlen_le by auto then show False using assms(2) by simp qed lemma canonical_take_canlen: assumes "bit_symbols zs" shows "canonical (take (canlen zs) zs)" proof (cases "canlen zs = 0") case True then show ?thesis using canonical_def by simp next case False then show ?thesis using canonical_def assms canlen_le_length canlen_one by (smt (verit, ccfv_SIG) One_nat_def Suc_pred append_take_drop_id diff_less last_length length_take less_le_trans min_absorb2 neq0_conv nth_append zero_less_one) qed lemma num_take_canlen_eq: "num (take (canlen zs) zs) = num zs" proof (induction "length zs - canlen zs" arbitrary: zs) case 0 then show ?case by simp next case (Suc x) let ?m = "canlen zs" have *: "\i ?m \ zs ! i = \" using canlen_at_ge by auto have "canlen zs < length zs" using Suc by simp then have "zs ! (length zs - 1) = \" using Suc canlen_at_ge canlen_le_length by (metis One_nat_def Suc_pred diff_less le_Suc_eq less_nat_zero_code nat_neq_iff zero_less_one) then have "todigit (zs ! (length zs - 1)) = 0" by simp moreover have ys: "zs = take (length zs - 1) zs @ [zs ! (length zs - 1)]" (is "zs = ?ys @ _") by (metis Suc_diff_1 \canlen zs < length zs\ append_butlast_last_id butlast_conv_take gr_implies_not0 last_length length_0_conv length_greater_0_conv) ultimately have "num ?ys = num zs" using num_trailing_zero by metis have canlen_ys: "canlen ?ys = canlen zs" proof (rule canlen_eqI) show "\i i \ ?ys ! i = \" by (simp add: canlen_at_ge) show "\y. \i i \ ?ys ! i = \ \ canlen zs \ y" using * Suc.hyps(2) canlen_le by (smt (verit, del_insts) One_nat_def Suc_pred append_take_drop_id diff_le_self length_take length_upt less_Suc_eq less_nat_zero_code list.size(3) min_absorb2 nth_append upt.simps(2) zero_less_Suc) qed then have "length ?ys - canlen ?ys = x" using ys Suc.hyps(2) by (metis butlast_snoc diff_Suc_1 diff_commute length_butlast) then have "num (take (canlen ?ys) ?ys) = num ?ys" using Suc by blast then have "num (take (canlen zs) ?ys) = num ?ys" using canlen_ys by simp then have "num (take (canlen zs) zs) = num ?ys" by (metis \canlen zs < length zs\ butlast_snoc take_butlast ys) then show ?case using \num ?ys = num zs\ by presburger qed lemma canrepr_take_canlen: assumes "num zs = n" and "bit_symbols zs" shows "canrepr n = take (canlen zs) zs" using assms canrepr canonical_canrepr canonical_ex1 canonical_take_canlen num_take_canlen_eq by blast lemma length_canrepr_canlen: assumes "num zs = n" and "bit_symbols zs" shows "nlength n = canlen zs" using canrepr_take_canlen assms canlen_le_length by (metis length_take min_absorb2) lemma nlength_ge_pow: assumes "nlength n = Suc j" shows "n \ 2 ^ j" proof - let ?xs = "canrepr n" have "?xs ! (length ?xs - 1) = \" using canonical_def assms canonical_canrepr by (metis Suc_neq_Zero diff_Suc_1 last_length length_0_conv) moreover have "(\i\[0.. todigit (?xs ! (length ?xs - 1)) * 2 ^ (length ?xs - 1)" using assms by simp ultimately have "num ?xs \ 2 ^ (length ?xs - 1)" using num_def by simp moreover have "num ?xs = n" using canrepr by simp ultimately show ?thesis using assms by simp qed lemma nlength_less_pow: "n < 2 ^ (nlength n)" proof (induction "nlength n" arbitrary: n) case 0 then show ?case by (metis canrepr canrepr_0 length_0_conv nat_zero_less_power_iff) next case (Suc j) let ?xs = "canrepr n" have lenxs: "length ?xs = Suc j" using Suc by simp have hdtl: "?xs = hd ?xs # tl ?xs" using Suc by (metis hd_Cons_tl list.size(3) nat.simps(3)) have len: "length (tl ?xs) = j" using Suc by simp have can: "canonical (tl ?xs)" using hdtl canonical_canrepr canonical_tl by metis define n' where "n' = num (tl ?xs)" then have "nlength n' = j" using len can canreprI by simp then have n'_less: "n' < 2 ^ j" using Suc by auto have "num ?xs = todigit (hd ?xs) + 2 * num (tl ?xs)" by (metis hdtl num_Cons) then have "n = todigit (hd ?xs) + 2 * num (tl ?xs)" using canrepr by simp also have "... \ 1 + 2 * num (tl ?xs)" by simp also have "... = 1 + 2 * n'" using n'_def by simp also have "... \ 1 + 2 * (2 ^ j - 1)" using n'_less by simp also have "... = 2 ^ (Suc j) - 1" by (metis (no_types, lifting) add_Suc_right le_add_diff_inverse mult_2 one_le_numeral one_le_power plus_1_eq_Suc sum.op_ivl_Suc sum_power2 zero_order(3)) also have "... < 2 ^ (Suc j)" by simp also have "... = 2 ^ (nlength n)" using lenxs by simp finally show ?case . qed lemma pow_nlength: assumes "2 ^ j \ n" and "n < 2 ^ (Suc j)" shows "nlength n = Suc j" proof (rule ccontr) assume "nlength n \ Suc j" then have "nlength n < Suc j \ nlength n > Suc j" by auto then show False proof assume "nlength n < Suc j" then have "nlength n \ j" by simp moreover have "n < 2 ^ (nlength n)" using nlength_less_pow by simp ultimately have "n < 2 ^ j" by (metis le_less_trans nat_power_less_imp_less not_less numeral_2_eq_2 zero_less_Suc) then show False using assms(1) by simp next assume *: "nlength n > Suc j" then have "n \ 2 ^ (nlength n - 1)" using nlength_ge_pow by simp moreover have "nlength n - 1 \ Suc j" using * by simp ultimately have "n \ 2 ^ (Suc j)" by (metis One_nat_def le_less_trans less_2_cases_iff linorder_not_less power_less_imp_less_exp) then show False using assms(2) by simp qed qed lemma nlength_le_n: "nlength n \ n" proof (cases "n = 0") case True then show ?thesis using canrepr_0 by simp next case False then have "nlength n > 0" using nlength_0 by simp moreover from this have "n \ 2 ^ (nlength n - 1)" using nlength_0 nlength_ge_pow by auto ultimately show ?thesis using nlength_ge_pow by (metis Suc_diff_1 Suc_leI dual_order.trans less_exp) qed lemma nlength_Suc_le: "nlength n \ nlength (Suc n)" proof (cases "n = 0") case True then show ?thesis by (simp add: canrepr_0) next case False then obtain j where j: "nlength n = Suc j" by (metis canrepr canrepr_0 gr0_implies_Suc length_greater_0_conv) then have "n \ 2 ^ j" using nlength_ge_pow by simp show ?thesis proof (cases "Suc n \ 2 ^ (Suc j)") case True have "n < 2 ^ (Suc j)" using j nlength_less_pow by metis then have "Suc n < 2 ^ (Suc (Suc j))" by simp then have "nlength (Suc n) = Suc (Suc j)" using True pow_nlength by simp then show ?thesis using j by simp next case False then have "Suc n < 2 ^ (Suc j)" by simp then have "nlength (Suc n) = Suc j" using `n \ 2 ^ j` pow_nlength by simp then show ?thesis using j by simp qed qed lemma nlength_mono: assumes "n1 \ n2" shows "nlength n1 \ nlength n2" proof - have "nlength n \ nlength (n + d)" for n d proof (induction d) case 0 then show ?case by simp next case (Suc d) then show ?case using nlength_Suc_le by (metis nat_arith.suc1 order_trans) qed then show ?thesis using assms by (metis le_add_diff_inverse) qed lemma nlength_even_le: "n > 0 \ nlength (2 * n) = Suc (nlength n)" proof - assume "n > 0" then have "nlength n > 0" by (metis canrepr canrepr_0 length_greater_0_conv less_numeral_extra(3)) then have "n \ 2 ^ (nlength n - 1)" using Suc_diff_1 nlength_ge_pow by simp then have "2 * n \ 2 ^ (nlength n)" by (metis Suc_diff_1 \0 < nlength n\ mult_le_mono2 power_Suc) moreover have "2 * n < 2 ^ (Suc (nlength n))" using nlength_less_pow by simp ultimately show ?thesis using pow_nlength by simp qed lemma nlength_prod: "nlength (n1 * n2) \ nlength n1 + nlength n2" proof - let ?j1 = "nlength n1" and ?j2 = "nlength n2" have "n1 < 2 ^ ?j1" "n2 < 2 ^ ?j2" using nlength_less_pow by simp_all then have "n1 * n2 < 2 ^ ?j1 * 2 ^ ?j2" by (simp add: mult_strict_mono) then have "n1 * n2 < 2 ^ (?j1 + ?j2)" by (simp add: power_add) then have "n1 * n2 \ 2 ^ (?j1 + ?j2) - 1" by simp then have "nlength (n1 * n2) \ nlength (2 ^ (?j1 + ?j2) - 1)" using nlength_mono by simp then show "nlength (n1 * n2) \ ?j1 + ?j2" using nlength_pow_minus_1 by simp qed text \ In the following lemma @{const Suc} is needed because $n^0 = 1$. \ lemma nlength_pow: "nlength (n ^ d) \ Suc (d * nlength n)" proof (induction d) case 0 then show ?case by (metis less_or_eq_imp_le mult_not_zero nat_power_eq_Suc_0_iff nlength_pow2) next case (Suc d) have "nlength (n ^ Suc d) = nlength (n ^ d * n)" by (simp add: mult.commute) then have "nlength (n ^ Suc d) \ nlength (n ^ d) + nlength n" using nlength_prod by simp then show ?case using Suc by simp qed lemma nlength_sum: "nlength (n1 + n2) \ Suc (max (nlength n1) (nlength n2))" proof - let ?m = "max n1 n2" have "n1 + n2 \ 2 * ?m" by simp then have "nlength (n1 + n2) \ nlength (2 * ?m)" using nlength_mono by simp moreover have "nlength ?m = max (nlength n1) (nlength n2)" using nlength_mono by (metis max.absorb1 max.cobounded2 max_def) ultimately show ?thesis using nlength_even_le by (metis canrepr_0 le_SucI le_zero_eq list.size(3) max_nat.neutr_eq_iff not_gr_zero zero_eq_add_iff_both_eq_0) qed lemma nlength_Suc: "nlength (Suc n) \ Suc (nlength n)" using nlength_sum nlength_1_simp by (metis One_nat_def Suc_leI add_Suc diff_Suc_1 length_greater_0_conv max.absorb_iff2 max.commute max_def nlength_0 plus_1_eq_Suc) lemma nlength_less_n: "n \ 3 \ nlength n < n" proof (induction n rule: nat_induct_at_least) case base then show ?case by (simp add: nlength_3) next case (Suc n) then show ?case using nlength_Suc by (metis Suc_le_eq le_neq_implies_less nlength_le_n not_less_eq) qed subsubsection \Comparing two numbers\ text \ In order to compare two numbers in canonical representation, we can use the Turing machine @{const tm_equals}, which works for arbitrary proper symbol sequences. \null \ lemma min_nlength: "min (nlength n1) (nlength n2) = nlength (min n1 n2)" by (metis min_absorb2 min_def nat_le_linear nlength_mono) lemma max_nlength: "max (nlength n1) (nlength n2) = nlength (max n1 n2)" using nlength_mono by (metis max.absorb1 max.cobounded2 max_def) lemma contents_blank_0: "\[\]\ = \[]\" using contents_def by auto definition tm_equalsn :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_equalsn \ tm_equals" lemma tm_equalsn_tm: assumes "k \ 2" and "G \ 4" and "0 < j3" shows "turing_machine k G (tm_equalsn j1 j2 j3)" unfolding tm_equalsn_def using assms tm_equals_tm by simp lemma transforms_tm_equalsnI [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and k b n1 n2 :: nat assumes "length tps = k" "j1 \ j2" "j2 \ j3" "j1 \ j3" "j1 < k" "j2 < k" "j3 < k" and "b \ 1" assumes "tps ! j1 = (\n1\\<^sub>N, 1)" "tps ! j2 = (\n2\\<^sub>N, 1)" "tps ! j3 = (\b\\<^sub>N, 1)" assumes "ttt = (3 * nlength (min n1 n2) + 7)" assumes "tps' = tps [j3 := (\if n1 = n2 then 1 else 0\\<^sub>N, 1)]" shows "transforms (tm_equalsn j1 j2 j3) tps ttt tps'" unfolding tm_equalsn_def proof (tform tps: assms) show "proper_symbols (canrepr n1)" using proper_symbols_canrepr by simp show "proper_symbols (canrepr n2)" using proper_symbols_canrepr by simp show "ttt = 3 * min (nlength n1) (nlength n2) + 7" using assms(12) min_nlength by simp let ?v = "if canrepr n1 = canrepr n2 then 3::nat else 0::nat" have "b = 0 \ b = 1" using assms(8) by auto then have "\b\\<^sub>N = \[]\ \ \b\\<^sub>N = \[\]\" using canrepr_0 canrepr_1 by auto then have "tps ! j3 = (\[]\, 1) \ tps ! j3 = (\[\]\, 1)" using assms(11) by simp then have v: "tps ! j3 |:=| ?v = (\[?v]\, 1)" using contents_def by auto show "tps' = tps[j3 := tps ! j3 |:=| ?v]" proof (cases "n1 = n2") case True then show ?thesis using canrepr_1 v assms(13) by auto next case False then have "?v = 0" by (metis canrepr) then show ?thesis using canrepr_0 v assms(13) contents_blank_0 by auto qed qed subsubsection \Copying a number between tapes\ text \ The next Turing machine overwrites the contents of tape $j_2$ with the contents of tape $j_1$ and performs a carriage return on both tapes. \ definition tm_copyn :: "tapeidx \ tapeidx \ machine" where "tm_copyn j1 j2 \ tm_erase_cr j2 ;; tm_cp_until j1 j2 {\} ;; tm_cr j1 ;; tm_cr j2" lemma tm_copyn_tm: assumes "k \ 2" and "G \ 4" and "j1 < k" "j2 < k" "j1 \ j2" "0 < j2" shows "turing_machine k G (tm_copyn j1 j2)" unfolding tm_copyn_def using assms tm_cp_until_tm tm_cr_tm tm_erase_cr_tm by simp locale turing_machine_move = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_erase_cr j2" definition "tm2 \ tm1 ;; tm_cp_until j1 j2 {\}" definition "tm3 \ tm2 ;; tm_cr j1" definition "tm4 \ tm3 ;; tm_cr j2" lemma tm4_eq_tm_copyn: "tm4 = tm_copyn j1 j2" unfolding tm4_def tm3_def tm2_def tm1_def tm_copyn_def by simp context fixes x y :: nat and tps0 :: "tape list" assumes j_less [simp]: "j1 < length tps0" "j2 < length tps0" assumes j [simp]: "j1 \ j2" and tps_j1 [simp]: "tps0 ! j1 = (\x\\<^sub>N, 1)" and tps_j2 [simp]: "tps0 ! j2 = (\y\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j2 := (\[]\, 1)]" lemma tm1 [transforms_intros]: assumes "t = 7 + 2 * nlength y" shows "transforms tm1 tps0 t tps1" unfolding tm1_def proof (tform tps: tps1_def time: assms) show "proper_symbols (canrepr y)" using proper_symbols_canrepr by simp qed definition "tps2 \ tps0 [j1 := (\x\\<^sub>N, Suc (nlength x)), j2 := (\x\\<^sub>N, Suc (nlength x))]" lemma tm2 [transforms_intros]: assumes "t = 8 + (2 * nlength y + nlength x)" shows "transforms tm2 tps0 t tps2" unfolding tm2_def proof (tform tps: tps1_def time: assms) show "rneigh (tps1 ! j1) {\} (nlength x)" proof (rule rneighI) show "(tps1 ::: j1) (tps1 :#: j1 + nlength x) \ {\}" using tps1_def canrepr_0 contents_outofbounds j(1) nlength_0_simp tps_j1 by (metis fst_eqD lessI nth_list_update_neq plus_1_eq_Suc singleton_iff snd_eqD) show "\n'. n' < nlength x \ (tps1 ::: j1) (tps1 :#: j1 + n') \ {\}" using tps1_def tps_j1 j j_less contents_inbounds proper_symbols_canrepr by (metis Suc_leI add_diff_cancel_left' fst_eqD not_add_less2 nth_list_update_neq plus_1_eq_Suc singletonD snd_eqD zero_less_Suc) qed have "(\x\\<^sub>N, Suc (nlength x)) = tps0[j2 := (\[]\, 1)] ! j1 |+| nlength x" using tps_j1 tps_j2 by (metis fst_eqD j(1) j_less(2) nth_list_update plus_1_eq_Suc snd_eqD) moreover have "(\x\\<^sub>N, Suc (nlength x)) = implant (tps0[j2 := (\[]\, 1)] ! j1) (tps0[j2 := (\[]\, 1)] ! j2) (nlength x)" using tps_j1 tps_j2 j j_less implant_contents nlength_0_simp by (metis add.right_neutral append.simps(1) canrepr_0 diff_Suc_1 drop0 le_eq_less_or_eq nth_list_update_eq nth_list_update_neq plus_1_eq_Suc take_all zero_less_one) ultimately show "tps2 = tps1 [j1 := tps1 ! j1 |+| nlength x, j2 := implant (tps1 ! j1) (tps1 ! j2) (nlength x)]" unfolding tps2_def tps1_def by (simp add: list_update_swap[of j1]) qed definition "tps3 \ tps0[j2 := (\x\\<^sub>N, Suc (nlength x))]" lemma tm3 [transforms_intros]: assumes "t = 11 + (2 * nlength y + 2 * nlength x)" shows "transforms tm3 tps0 t tps3" unfolding tm3_def proof (tform tps: tps2_def) have "tps2 :#: j1 = Suc (nlength x)" using assms tps2_def by (metis j(1) j_less(1) nth_list_update_eq nth_list_update_neq snd_conv) then show "t = 8 + (2 * nlength y + nlength x) + (tps2 :#: j1 + 2)" using assms by simp show "clean_tape (tps2 ! j1)" using tps2_def by (simp add: clean_tape_ncontents nth_list_update_neq') have "tps2 ! j1 |#=| 1 = (\x\\<^sub>N, 1)" using tps2_def by (simp add: nth_list_update_neq') then show "tps3 = tps2[j1 := tps2 ! j1 |#=| 1]" using tps3_def tps2_def by (metis j(1) list_update_id list_update_overwrite list_update_swap tps_j1) qed definition "tps4 \ tps0[j2 := (\x\\<^sub>N, 1)]" lemma tm4: assumes "t = 14 + (3 * nlength x + 2 * nlength y)" shows "transforms tm4 tps0 t tps4" unfolding tm4_def proof (tform tps: tps3_def time: tps3_def assms) show "clean_tape (tps3 ! j2)" using tps3_def clean_tape_ncontents by simp have "tps3 ! j2 |#=| 1 = (\x\\<^sub>N, 1)" using tps3_def by (simp add: nth_list_update_neq') then show "tps4 = tps3[j2 := tps3 ! j2 |#=| 1]" using tps4_def tps3_def by (metis list_update_overwrite tps_j1) qed lemma tm4': assumes "t = 14 + 3 * (nlength x + nlength y)" shows "transforms tm4 tps0 t tps4" using tm4 transforms_monotone assms by simp end end (* locale turing_machine_move *) lemma transforms_tm_copynI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k x y :: nat assumes "j1 \ j2" "j1 < length tps" "j2 < length tps" assumes "tps ! j1 = (\x\\<^sub>N, 1)" "tps ! j2 = (\y\\<^sub>N, 1)" assumes "ttt = 14 + 3 * (nlength x + nlength y)" assumes "tps' = tps [j2 := (\x\\<^sub>N, 1)]" shows "transforms (tm_copyn j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_move j1 j2 . show ?thesis using assms loc.tm4' loc.tps4_def loc.tm4_eq_tm_copyn by simp qed subsubsection \Setting the tape contents to a number\ text \ The Turing machine in this section writes a hard-coded number to a tape. \ definition tm_setn :: "tapeidx \ nat \ machine" where "tm_setn j n \ tm_set j (canrepr n)" lemma tm_setn_tm: assumes "k \ 2" and "G \ 4" and "j < k" and "0 < j " shows "turing_machine k G (tm_setn j n)" proof - have "symbols_lt G (canrepr n)" using assms(2) bit_symbols_canrepr by fastforce then show ?thesis unfolding tm_setn_def using tm_set_tm assms by simp qed lemma transforms_tm_setnI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and x k n :: nat assumes "j < length tps" assumes "tps ! j = (\x\\<^sub>N, 1)" assumes "t = 10 + 2 * nlength x + 2 * nlength n" assumes "tps' = tps[j := (\n\\<^sub>N, 1)]" shows "transforms (tm_setn j n) tps t tps'" unfolding tm_setn_def using transforms_tm_setI[OF assms(1), of "canrepr x" "canrepr n" t tps'] assms canonical_canrepr canonical_def contents_clean_tape' by (simp add: eval_nat_numeral(3) numeral_Bit0 proper_symbols_canrepr) subsection \Incrementing\ text \ In this section we devise a Turing machine that increments a number. The next function describes how the symbol sequence of the incremented number looks like. Basically one has to flip all @{text \} symbols starting at the least significant digit until one reaches a @{text \}, which is then replaced by a @{text \}. If there is no @{text \}, a @{text \} is appended. Here we exploit that the most significant digit is to the right. \ definition nincr :: "symbol list \ symbol list" where "nincr zs \ if \i then replicate (LEAST i. i < length zs \ zs ! i = \) \ @ [\] @ drop (Suc (LEAST i. i < length zs \ zs ! i = \)) zs else replicate (length zs) \ @ [\]" lemma canonical_nincr: assumes "canonical zs" shows "canonical (nincr zs)" proof - have 1: "bit_symbols zs" using canonical_def assms by simp let ?j = "LEAST i. i < length zs \ zs ! i = \" have "bit_symbols (nincr zs)" proof (cases "\i") case True then have "nincr zs = replicate ?j \ @ [\] @ drop (Suc ?j) zs" using nincr_def by simp moreover have "bit_symbols (replicate ?j \)" by simp moreover have "bit_symbols [\]" by simp moreover have "bit_symbols (drop (Suc ?j) zs)" using 1 by simp ultimately show ?thesis using bit_symbols_append by presburger next case False then show ?thesis using nincr_def bit_symbols_append by auto qed moreover have "last (nincr zs) = \" proof (cases "\i") case True then show ?thesis using nincr_def assms canonical_def by auto next case False then show ?thesis using nincr_def by auto qed ultimately show ?thesis using canonical_def by simp qed lemma nincr: assumes "bit_symbols zs" shows "num (nincr zs) = Suc (num zs)" proof (cases "\i") case True define j where "j = (LEAST i. i < length zs \ zs ! i = \)" then have 1: "j < length zs \ zs ! j = \" using LeastI_ex[OF True] by simp have 2: "zs ! i = \" if "i < j" for i using that True j_def assms "1" less_trans not_less_Least by blast define xs :: "symbol list" where "xs = replicate j \ @ [\]" define ys :: "symbol list" where "ys = drop (Suc j) zs" have "zs = xs @ ys" proof - have "xs = take (Suc j) zs" using xs_def 1 2 by (smt (verit, best) le_eq_less_or_eq length_replicate length_take min_absorb2 nth_equalityI nth_replicate nth_take take_Suc_conv_app_nth) then show ?thesis using ys_def by simp qed have "nincr zs = replicate j \ @ [\] @ drop (Suc j) zs" using nincr_def True j_def by simp then have "num (nincr zs) = num (replicate j \ @ [\] @ ys)" using ys_def by simp also have "... = num (replicate j \ @ [\]) + 2 ^ Suc j * num ys" using num_append by (metis append_assoc length_append_singleton length_replicate) also have "... = Suc (num xs) + 2 ^ Suc j * num ys" proof - have "num (replicate j \ @ [\]) = 2 ^ j" using num_replicate2_eq_pow by simp also have "... = Suc (2 ^ j - 1)" by simp also have "... = Suc (num (replicate j \))" using num_replicate3_eq_pow_minus_1 by simp also have "... = Suc (num (replicate j \ @ [\]))" using num_trailing_zero by simp finally have "num (replicate j \ @ [\]) = Suc (num xs)" using xs_def by simp then show ?thesis by simp qed also have "... = Suc (num xs + 2 ^ Suc j * num ys)" by simp also have "... = Suc (num zs)" using `zs = xs @ ys` num_append xs_def by (metis length_append_singleton length_replicate) finally show ?thesis . next case False then have "\i" using assms by simp then have zs: "zs = replicate (length zs) \" by (simp add: nth_equalityI) then have num_zs: "num zs = 2 ^ length zs - 1" by (metis num_replicate3_eq_pow_minus_1) have "nincr zs = replicate (length zs) \ @ [\]" using nincr_def False by auto then have "num (nincr zs) = 2 ^ length zs" by (simp add: num_replicate2_eq_pow) then show ?thesis using num_zs by simp qed lemma nincr_canrepr: "nincr (canrepr n) = canrepr (Suc n)" using canrepr canonical_canrepr canreprI bit_symbols_canrepr canonical_nincr nincr by metis text \ The next Turing machine performs the incrementing. Starting from the left of the symbol sequence on tape $j$, it writes the symbol \textbf{0} until it reaches a blank or the symbol \textbf{1}. Then it writes a \textbf{1} and returns the tape head to the beginning. \ definition tm_incr :: "tapeidx \ machine" where "tm_incr j \ tm_const_until j j {\, \} \ ;; tm_write j \ ;; tm_cr j" lemma tm_incr_tm: assumes "G \ 4" and "k \ 2" and "j < k" and "j > 0" shows "turing_machine k G (tm_incr j)" unfolding tm_incr_def using assms tm_const_until_tm tm_write_tm tm_cr_tm by simp locale turing_machine_incr = fixes j :: tapeidx begin definition "tm1 \ tm_const_until j j {\, \} \" definition "tm2 \ tm1 ;; tm_write j \" definition "tm3 \ tm2 ;; tm_cr j" lemma tm3_eq_tm_incr: "tm3 = tm_incr j" unfolding tm3_def tm2_def tm1_def tm_incr_def by simp context fixes x k :: nat and tps :: "tape list" assumes jk [simp]: "j < k" "length tps = k" and tps0 [simp]: "tps ! j = (\x\\<^sub>N, 1)" begin lemma tm1 [transforms_intros]: assumes "i0 = (LEAST i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \})" and "tps' = tps[j := constplant (tps ! j) \ i0]" shows "transforms tm1 tps (Suc i0) tps'" unfolding tm1_def proof (tform tps: assms(2)) let ?P = "\i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \}" have 2: "i0 \ nlength x \ \x\\<^sub>N (Suc i0) \ {\, \}" using LeastI[of ?P "nlength x"] jk(1) assms(1) by simp have 3: "\ ?P i" if "i < i0" for i using not_less_Least[of i ?P] jk(1) assms(1) that by simp show "rneigh (tps ! j) {\, \} i0" proof (rule rneighI) show "(tps ::: j) (tps :#: j + i0) \ {\, \}" using tps0 2 jk(1) assms(1) by simp show "\n'. n' < i0 \ (tps ::: j) (tps :#: j + n') \ {\, \}" using tps0 2 3 jk(1) assms(1) by simp qed qed lemma tm2 [transforms_intros]: assumes "i0 = (LEAST i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \})" and "ttt = Suc (Suc i0)" and "tps' = tps[j := (\Suc x\\<^sub>N, Suc i0)]" shows "transforms tm2 tps ttt tps'" unfolding tm2_def proof (tform tps: assms(1,3) time: assms(1,2)) let ?P = "\i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \}" have 1: "?P (nlength x)" by simp have 2: "i0 \ nlength x \ \x\\<^sub>N (Suc i0) \ {\, \}" using LeastI[of ?P "nlength x"] assms(1) by simp have 3: "\ ?P i" if "i < i0" for i using not_less_Least[of i ?P] assms(1) that by simp let ?i = "LEAST i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \}" show "tps' = tps [j := constplant (tps ! j) 2 ?i, j := tps[j := constplant (tps ! j) \ ?i] ! j |:=| \]" (is "tps' = ?rhs") proof - have "?rhs = tps [j := constplant (\x\\<^sub>N, Suc 0) \ i0 |:=| \]" using jk assms(1) by simp moreover have "(\Suc x\\<^sub>N, Suc i0) = constplant (\x\\<^sub>N, Suc 0) 2 i0 |:=| \" (is "?l = ?r") proof - have "snd ?l = snd ?r" by (simp add: transplant_def) moreover have "\Suc x\\<^sub>N = fst ?r" proof - let ?zs = "canrepr x" have l: "\Suc x\\<^sub>N = \nincr ?zs\" by (simp add: nincr_canrepr) have r: "fst ?r = (\i. if Suc 0 \ i \ i < Suc i0 then \ else \x\\<^sub>N i)(Suc i0 := \)" using constplant by auto show ?thesis proof (cases "\i") case True let ?Q = "\i. i < length ?zs \ ?zs ! i = \" have Q1: "?Q (Least ?Q)" using True by (metis (mono_tags, lifting) LeastI_ex) have Q2: "\ ?Q i" if "i < Least ?Q" for i using True not_less_Least that by blast have "Least ?P = Least ?Q" proof (rule Least_equality) show "Least ?Q \ nlength x \ \x\\<^sub>N (Suc (Least ?Q)) \ {\, \}" proof show "Least ?Q \ nlength x" using True by (metis (mono_tags, lifting) LeastI_ex less_imp_le) show "\x\\<^sub>N (Suc (Least ?Q)) \ {\, \}" using True by (simp add: Q1 Suc_leI) qed then show "\y. y \ nlength x \ \x\\<^sub>N (Suc y) \ {\, \} \ (Least ?Q) \ y" using True Q1 Q2 bit_symbols_canrepr contents_def - by (smt (z3) Least_le Suc_leI bot_nat_0.not_eq_extremum diff_Suc_1 insert_iff le_neq_implies_less - nat.simps(3) nlength_0_simp nlength_le_n nlength_less_n singletonD) + by (metis (no_types, lifting) Least_le antisym_conv2 diff_Suc_1 insert_iff + le_less_Suc_eq less_nat_zero_code nat_le_linear proper_symbols_canrepr singletonD) qed then have i0: "i0 = Least ?Q" using assms(1) by simp then have nincr_zs: "nincr ?zs = replicate i0 \ @ [\] @ drop (Suc i0) ?zs" using nincr_def True by simp show ?thesis proof fix i consider "i = 0" | "Suc 0 \ i \ i < Suc i0" | "i = Suc i0" | "i > Suc i0 \ i \ length ?zs" | "i > Suc i0 \ i > length ?zs" by linarith then have "\replicate i0 \ @ [\] @ drop (Suc i0) ?zs\ i = ((\i. if Suc 0 \ i \ i < Suc i0 then \ else \x\\<^sub>N i)(Suc i0 := \)) i" (is "?A i = ?B i") proof (cases) case 1 then show ?thesis by (simp add: transplant_def) next case 2 then have "i - 1 < i0" by auto then have "(replicate i0 \ @ [\] @ drop (Suc i0) ?zs) ! (i - 1) = \" by (metis length_replicate nth_append nth_replicate) then have "?A i = \" using contents_def i0 "2" Q1 nincr_canrepr nincr_zs by (metis Suc_le_lessD le_trans less_Suc_eq_le less_imp_le_nat less_numeral_extra(3) nlength_Suc_le) moreover have "?B i = \" using i0 2 by simp ultimately show ?thesis by simp next case 3 then show ?thesis using i0 Q1 canrepr_0 contents_inbounds nincr_canrepr nincr_zs nlength_0_simp nlength_Suc nlength_Suc_le - by (smt (z3) Suc_leI append_Cons diff_Suc_1 fun_upd_apply le_trans length_replicate - nth_append_length zero_less_Suc) + by (metis (no_types, lifting) contents_append_update fun_upd_apply length_replicate) next case 4 then have "?A i = (replicate i0 \ @ [\] @ drop (Suc i0) ?zs) ! (i - 1)" by auto then have "?A i = ((replicate i0 \ @ [\]) @ drop (Suc i0) ?zs) ! (i - 1)" by simp moreover have "length (replicate i0 \ @ [\]) = Suc i0" by simp moreover have "i - 1 < length ?zs" using 4 by auto moreover have "i - 1 >= Suc i0" using 4 by auto ultimately have "?A i = ?zs ! (i - 1)" using i0 Q1 by (metis (no_types, lifting) Suc_leI append_take_drop_id length_take min_absorb2 not_le nth_append) moreover have "?B i = \x\\<^sub>N i" using 4 by simp ultimately show ?thesis using i0 4 contents_def by simp next case 5 then show ?thesis by auto qed then show "\Suc x\\<^sub>N i = fst (constplant (\x\\<^sub>N, Suc 0) \ i0 |:=| \) i" using nincr_zs l r by simp qed next case False then have nincr_zs: "nincr ?zs = replicate (length ?zs) \ @ [\]" using nincr_def by auto have "Least ?P = length ?zs" proof (rule Least_equality) show "nlength x \ nlength x \ \x\\<^sub>N (Suc (nlength x)) \ {\, \}" by simp show "\y. y \ nlength x \ \x\\<^sub>N (Suc y) \ {\, \} \ nlength x \ y" using False contents_def bit_symbols_canrepr by (metis diff_Suc_1 insert_iff le_neq_implies_less nat.simps(3) not_less_eq_eq numeral_3_eq_3 singletonD) qed then have i0: "i0 = length ?zs" using assms(1) by simp show ?thesis proof fix i consider "i = 0" | "Suc 0 \ i \ i < Suc (length ?zs)" | "i = Suc (length ?zs)" | "i > Suc (length ?zs)" by linarith then have "\replicate (length ?zs) \ @ [\]\ i = ((\i. if Suc 0 \ i \ i < Suc i0 then \ else \x\\<^sub>N i)(Suc i0 := \)) i" (is "?A i = ?B i") proof (cases) case 1 then show ?thesis by (simp add: transplant_def) next case 2 then have "?A i = \" by (metis One_nat_def Suc_le_lessD add.commute contents_def diff_Suc_1 length_Cons length_append length_replicate less_Suc_eq_0_disj less_imp_le_nat less_numeral_extra(3) list.size(3) nth_append nth_replicate plus_1_eq_Suc) moreover have "?B i = \" using i0 2 by simp ultimately show ?thesis by simp next case 3 then show ?thesis using i0 canrepr_0 contents_inbounds nincr_canrepr nincr_zs nlength_0_simp nlength_Suc by (metis One_nat_def add.commute diff_Suc_1 fun_upd_apply length_Cons length_append length_replicate nth_append_length plus_1_eq_Suc zero_less_Suc) next case 4 then show ?thesis using i0 by simp qed then show "\Suc x\\<^sub>N i = fst (constplant (\x\\<^sub>N, Suc 0) \ i0 |:=| \) i" using nincr_zs l r by simp qed qed qed ultimately show ?thesis by simp qed ultimately show ?thesis using assms(3) by simp qed qed lemma tm3: assumes "i0 = (LEAST i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \})" and "ttt = 5 + 2 * i0" and "tps' = tps[j := (\Suc x\\<^sub>N, Suc 0)]" shows "transforms tm3 tps ttt tps'" unfolding tm3_def proof (tform tps: assms(1,3) time: assms(1,2)) let ?tps = "tps[j := (\Suc x\\<^sub>N, Suc (LEAST i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \}))]" show "clean_tape (?tps ! j)" using clean_tape_ncontents by (simp add: assms(1,3)) qed lemma tm3': assumes "ttt = 5 + 2 * nlength x" and "tps' = tps[j := (\Suc x\\<^sub>N, Suc 0)]" shows "transforms tm3 tps ttt tps'" proof - let ?P = "\i. i \ nlength x \ \x\\<^sub>N (Suc i) \ {\, \}" define i0 where "i0 = Least ?P" have "i0 \ nlength x \ \x\\<^sub>N (Suc i0) \ {\, \}" using LeastI[of ?P "nlength x"] i0_def by simp then have "5 + 2 * i0 \ 5 + 2 * nlength x" by simp moreover have "transforms tm3 tps (5 + 2 * i0) tps'" using assms tm3 i0_def by simp ultimately show ?thesis using transforms_monotone assms(1) by simp qed end (* context *) end (* locale *) lemma transforms_tm_incrI [transforms_intros]: assumes "j < k" and "length tps = k" and "tps ! j = (\x\\<^sub>N, 1)" and "ttt = 5 + 2 * nlength x" and "tps' = tps[j := (\Suc x\\<^sub>N, 1)]" shows "transforms (tm_incr j) tps ttt tps'" proof - interpret loc: turing_machine_incr j . show ?thesis using assms loc.tm3' loc.tm3_eq_tm_incr by simp qed subsubsection \Incrementing multiple times\ text \ Adding a constant by iteratively incrementing is not exactly efficient, but it still only takes constant time and thus does not endanger any time bounds. \ fun tm_plus_const :: "nat \ tapeidx \ machine" where "tm_plus_const 0 j = []" | "tm_plus_const (Suc c) j = tm_plus_const c j ;; tm_incr j" lemma tm_plus_const_tm: assumes "k \ 2" and "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_plus_const c j)" using assms Nil_tm tm_incr_tm by (induction c) simp_all lemma transforms_tm_plus_constI [transforms_intros]: fixes c :: nat assumes "j < k" and "j > 0" and "length tps = k" and "tps ! j = (\x\\<^sub>N, 1)" and "ttt = c * (5 + 2 * nlength (x + c))" and "tps' = tps[j := (\x + c\\<^sub>N, 1)]" shows "transforms (tm_plus_const c j) tps ttt tps'" using assms(5,6,4) proof (induction c arbitrary: ttt tps') case 0 then show ?case using transforms_Nil assms by (metis add_cancel_left_right list_update_id mult_eq_0_iff tm_plus_const.simps(1)) next case (Suc c) define tpsA where "tpsA = tps[j := (\x + c\\<^sub>N, 1)]" let ?ttt = "c * (5 + 2 * nlength (x + c)) + (5 + 2 * nlength (x + c))" have "transforms (tm_plus_const c j ;; tm_incr j) tps ?ttt tps'" proof (tform tps: assms) show "transforms (tm_plus_const c j) tps (c * (5 + 2 * nlength (x + c))) tpsA" using tpsA_def assms Suc by simp show "j < length tpsA" using tpsA_def assms(1,3) by simp show "tpsA ! j = (\x + c\\<^sub>N, 1)" using tpsA_def assms(1,3) by simp show "tps' = tpsA[j := (\Suc (x + c)\\<^sub>N, 1)]" using tpsA_def assms Suc by (metis add_Suc_right list_update_overwrite) qed moreover have "?ttt \ ttt" proof - have "?ttt = Suc c * (5 + 2 * nlength (x + c))" by simp also have "... \ Suc c * (5 + 2 * nlength (x + Suc c))" using nlength_mono Suc_mult_le_cancel1 by auto finally show "?ttt \ ttt" using Suc by simp qed ultimately have "transforms (tm_plus_const c j ;; tm_incr j) tps ttt tps'" using transforms_monotone by simp then show ?case by simp qed subsection \Decrementing\ text \ Decrementing a number is almost like incrementing but with the symbols \textbf{0} and \textbf{1} swapped. One difference is that in order to get a canonical symbol sequence, a trailing zero must be removed, whereas incrementing cannot result in a trailing zero. Another difference is that decrementing the number zero yields zero. The next function returns the leftmost symbol~\textbf{1}, that is, the one that needs to be flipped. \ definition first1 :: "symbol list \ nat" where "first1 zs \ LEAST i. i < length zs \ zs ! i = \" lemma canonical_ex_3: assumes "canonical zs" and "zs \ []" shows "\i" using assms canonical_def by (metis One_nat_def Suc_pred last_conv_nth length_greater_0_conv lessI) lemma canonical_first1: assumes "canonical zs" and "zs \ []" shows "first1 zs < length zs \ zs ! first1 zs = \" using assms canonical_ex_3 by (metis (mono_tags, lifting) LeastI first1_def) lemma canonical_first1_less: assumes "canonical zs" and "zs \ []" shows "\i" proof - have "\i \" using assms first1_def canonical_first1 not_less_Least by fastforce then show ?thesis using assms canonical_def by (meson canonical_first1 less_trans) qed text \ The next function describes how the canonical representation of the decremented symbol sequence looks like. It has special cases for the empty sequence and for sequences whose only \textbf{1} is the most significant digit. \ definition ndecr :: "symbol list \ symbol list" where "ndecr zs \ if zs = [] then [] else if first1 zs = length zs - 1 then replicate (first1 zs) \ else replicate (first1 zs) \ @ [\] @ drop (Suc (first1 zs)) zs" lemma canonical_ndecr: assumes "canonical zs" shows "canonical (ndecr zs)" proof - let ?i = "first1 zs" consider "zs = []" | "zs \ [] \ first1 zs = length zs - 1" | "zs \ [] \ first1 zs < length zs - 1" using canonical_first1 assms by fastforce then show ?thesis proof (cases) case 1 then show ?thesis using ndecr_def canonical_def by simp next case 2 then show ?thesis using canonical_def ndecr_def not_less_eq by fastforce next case 3 then have "Suc (first1 zs) < length zs" by auto then have "last (drop (Suc (first1 zs)) zs) = \" using assms canonical_def 3 by simp moreover have "bit_symbols (replicate (first1 zs) \ @ [\] @ drop (Suc (first1 zs)) zs)" proof - have "bit_symbols (replicate (first1 zs) \)" by simp moreover have "bit_symbols [\]" by simp moreover have "bit_symbols (drop (Suc (first1 zs)) zs)" using assms canonical_def by simp ultimately show ?thesis using bit_symbols_append by presburger qed ultimately show ?thesis using canonical_def ndecr_def 3 by auto qed qed lemma ndecr: assumes "canonical zs" shows "num (ndecr zs) = num zs - 1" proof - let ?i = "first1 zs" consider "zs = []" | "zs \ [] \ first1 zs = length zs - 1" | "zs \ [] \ first1 zs < length zs - 1" using canonical_first1 assms by fastforce then show ?thesis proof (cases) case 1 then show ?thesis using ndecr_def canrepr_0 canrepr by (metis zero_diff) next case 2 then have less: "zs ! i = \" if "i < first1 zs" for i using that assms canonical_first1_less by simp have at: "zs ! (first1 zs) = \" using 2 canonical_first1 assms by blast have "zs = replicate (first1 zs) \ @ [\]" (is "zs = ?zs") proof (rule nth_equalityI) show len: "length zs = length ?zs" using 2 by simp show "zs ! i = ?zs ! i" if "i < length zs" for i proof (cases "i < first1 zs") case True then show ?thesis by (simp add: less nth_append) next case False then show ?thesis using len that at by (metis Suc_leI leD length_append_singleton length_replicate linorder_neqE_nat nth_append_length) qed qed moreover from this have "ndecr zs = replicate (first1 zs) 3" using ndecr_def 2 by simp ultimately show ?thesis using num_replicate2_eq_pow num_replicate3_eq_pow_minus_1 by metis next case 3 then have less: "zs ! i = \" if "i < ?i" for i using that assms canonical_first1_less by simp have at: "zs ! ?i = \" using 3 canonical_first1 assms by simp have zs: "zs = replicate ?i \ @ [\] @ drop (Suc ?i) zs" (is "zs = ?zs") proof (rule nth_equalityI) show len: "length zs = length ?zs" using 3 by auto show "zs ! i = ?zs ! i" if "i < length zs" for i proof - consider "i < ?i" | "i = ?i" | "i > ?i" by linarith then show ?thesis proof (cases) case 1 then show ?thesis using less by (metis length_replicate nth_append nth_replicate) next case 2 then show ?thesis using at by (metis append_Cons length_replicate nth_append_length) next case 3 have "?zs = (replicate ?i \ @ [\]) @ drop (Suc ?i) zs" by simp then have "?zs ! i = drop (Suc ?i) zs ! (i - Suc ?i)" using 3 by (simp add: nth_append) then have "?zs ! i = zs ! i" using 3 that by simp then show ?thesis by simp qed qed qed then have "ndecr zs = replicate ?i \ @ [\] @ drop (Suc ?i) zs" using ndecr_def 3 by simp then have "Suc (num (ndecr zs)) = Suc (num ((replicate ?i \ @ [\]) @ drop (Suc ?i) zs))" (is "_ = Suc (num (?xs @ ?ys))") by simp also have "... = Suc (num ?xs + 2 ^ length ?xs * num ?ys)" using num_append by blast also have "... = Suc (num ?xs + 2 ^ Suc ?i * num ?ys)" by simp also have "... = Suc (2 ^ ?i - 1 + 2 ^ Suc ?i * num ?ys)" using num_replicate3_eq_pow_minus_1 num_trailing_zero[of 2 "replicate ?i \"] by simp also have "... = 2 ^ ?i + 2 ^ Suc ?i * num ?ys" by simp also have "... = num (replicate ?i \ @ [\]) + 2 ^ Suc ?i * num ?ys" using num_replicate2_eq_pow by simp also have "... = num ((replicate ?i \ @ [\]) @ ?ys)" using num_append by (metis length_append_singleton length_replicate) also have "... = num (replicate ?i \ @ [\] @ ?ys)" by simp also have "... = num zs" using zs by simp finally have "Suc (num (ndecr zs)) = num zs" . then show ?thesis by simp qed qed text \ The next Turing machine implements the function @{const ndecr}. It does nothing on the empty input, which represents zero. On other inputs it writes symbols \textbf{1} going right until it reaches a \textbf{1} symbol, which is guaranteed to happen for non-empty canonical representations. It then overwrites this \textbf{1} with \textbf{0}. If there is a blank symbol to the right of this \textbf{0}, the \textbf{0} is removed again. \ definition tm_decr :: "tapeidx \ machine" where "tm_decr j \ IF \rs. rs ! j = \ THEN [] ELSE tm_const_until j j {\} \ ;; tm_rtrans j (\_. \) ;; IF \rs. rs ! j = \ THEN tm_left j ;; tm_write j \ ELSE [] ENDIF ;; tm_cr j ENDIF" lemma tm_decr_tm: assumes "G \ 4" and "k \ 2" and "j < k" and "0 < j" shows "turing_machine k G (tm_decr j)" unfolding tm_decr_def using assms tm_cr_tm tm_const_until_tm tm_rtrans_tm tm_left_tm tm_write_tm turing_machine_branch_turing_machine Nil_tm by simp locale turing_machine_decr = fixes j :: tapeidx begin definition "tm1 \ tm_const_until j j {\} \" definition "tm2 \ tm1 ;; tm_rtrans j (\_. \)" definition "tm23 \ tm_left j" definition "tm24 \ tm23 ;; tm_write j \" definition "tm25 \ IF \rs. rs ! j = \ THEN tm24 ELSE [] ENDIF" definition "tm5 \ tm2 ;; tm25" definition "tm6 \ tm5 ;; tm_cr j" definition "tm7 \ IF \rs. rs ! j = \ THEN [] ELSE tm6 ENDIF" lemma tm7_eq_tm_decr: "tm7 = tm_decr j" unfolding tm1_def tm2_def tm23_def tm24_def tm25_def tm5_def tm6_def tm7_def tm_decr_def by simp context fixes tps0 :: "tape list" and xs :: "symbol list" and k :: nat assumes jk: "length tps0 = k" "j < k" and can: "canonical xs" and tps0: "tps0 ! j = (\xs\, 1)" begin lemma bs: "bit_symbols xs" using can canonical_def by simp context assumes read_tps0: "read tps0 ! j = \" begin lemma xs_Nil: "xs = []" using tps0 jk tapes_at_read' read_tps0 bs contents_inbounds by (metis can canreprI canrepr_0 fst_conv ncontents_1_blank_iff_zero snd_conv) lemma transforms_NilI: assumes "ttt = 0" and "tps' = tps0[j := (\ndecr xs\, 1)]" shows "transforms [] tps0 ttt tps'" using transforms_Nil xs_Nil ndecr_def tps0 assms by (metis Basics.transforms_Nil list_update_id) end (* context read tps0 ! j = 0 *) context assumes read_tps0': "read tps0 ! j \ \" begin lemma xs: "xs \ []" using tps0 jk tapes_at_read' read_tps0' bs contents_inbounds by (metis canrepr_0 fst_conv ncontents_1_blank_iff_zero snd_conv) lemma first1: "first1 xs < length xs" "xs ! first1 xs = \" "\i" using canonical_first1[OF can xs] canonical_first1_less[OF can xs] by simp_all definition "tps1 \ tps0 [j := (\replicate (first1 xs) \ @ [\] @ (drop (Suc (first1 xs)) xs)\, Suc (first1 xs))]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (first1 xs)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps1_def jk time: assms) show "rneigh (tps0 ! j) {\} (first1 xs)" proof (rule rneighI) show "(tps0 ::: j) (tps0 :#: j + first1 xs) \ {\}" using first1(1,2) tps0 jk by (simp add: Suc_leI) show "\n'. n' < first1 xs \ (tps0 ::: j) (tps0 :#: j + n') \ {\}" using first1(3) tps0 jk by (simp add: contents_def) qed show "tps1 = tps0 [j := tps0 ! j |+| first1 xs, j := constplant (tps0 ! j) \ (first1 xs)]" proof - have "tps1 ! j = constplant (tps0 ! j) 3 (first1 xs)" (is "_ = ?rhs") proof - have "fst ?rhs = (\i. if 1 \ i \ i < 1 + first1 xs then \ else \xs\ i)" using tps0 jk constplant by auto also have "... = \replicate (first1 xs) \ @ [\] @ drop (Suc (first1 xs)) xs\" proof fix i consider "i = 0" | "i \ 1 \ i < 1 + first1 xs" | "i = 1 + first1 xs" | "1 + first1 xs < i \ i \ length xs" | "i > length xs" by linarith then show "(if 1 \ i \ i < 1 + first1 xs then \ else \xs\ i) = \replicate (first1 xs) \ @ [\] @ drop (Suc (first1 xs)) xs\ i" (is "?l = ?r") proof (cases) case 1 then show ?thesis by simp next case 2 then show ?thesis by (smt (verit) One_nat_def Suc_diff_Suc add_diff_inverse_nat contents_inbounds first1(1) length_append length_drop length_replicate less_imp_le_nat less_le_trans list.size(3) list.size(4) not_le not_less_eq nth_append nth_replicate plus_1_eq_Suc) next case 3 then show ?thesis using first1 by (smt (verit) One_nat_def Suc_diff_Suc Suc_leI add_diff_inverse_nat append_Cons contents_inbounds diff_Suc_1 length_append length_drop length_replicate less_SucI less_Suc_eq_0_disj list.size(3) list.size(4) not_less_eq nth_append_length) next case 4 then have "?r = (replicate (first1 xs) \ @ [\] @ drop (Suc (first1 xs)) xs) ! (i - 1)" by auto also have "... = ((replicate (first1 xs) \ @ [\]) @ drop (Suc (first1 xs)) xs) ! (i - 1)" by simp also have "... = (drop (Suc (first1 xs)) xs) ! (i - 1 - Suc (first1 xs))" using 4 by (metis Suc_leI add_diff_inverse_nat gr_implies_not0 leD length_append_singleton length_replicate less_one nth_append plus_1_eq_Suc) also have "... = xs ! (i - 1)" using 4 by (metis Suc_leI add_diff_inverse_nat first1(1) gr_implies_not0 leD less_one nth_drop plus_1_eq_Suc) also have "... = \xs\ i" using 4 by simp also have "... = ?l" using 4 by simp finally have "?r = ?l" . then show ?thesis by simp next case 5 then show ?thesis using first1(1) by simp qed qed also have "... = tps1 ::: j" using tps1_def jk by simp finally have "fst ?rhs = fst (tps1 ! j)" . then show ?thesis using tps1_def jk constplant tps0 by simp qed then show ?thesis using tps1_def tps0 jk by simp qed qed definition "tps2 \ tps0 [j := (\replicate (first1 xs) \ @ [\] @ drop (Suc (first1 xs)) xs\, Suc (Suc (first1 xs)))]" lemma tm2 [transforms_intros]: assumes "ttt = first1 xs + 2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps2_def tps1_def jk time: assms) show "tps2 = tps1[j := tps1 ! j |:=| \ |+| 1]" using tps1_def tps2_def jk contents_append_update by simp qed definition "tps5 \ tps0 [j := (\ndecr xs\, if read tps2 ! j = \ then Suc (first1 xs) else Suc (Suc (first1 xs)))]" context assumes read_tps2: "read tps2 ! j = \" begin lemma proper_contents_outofbounds: assumes "proper_symbols zs" and "\zs\ i = \" shows "i > length zs" using contents_def proper_symbols_ne0 assms by (metis Suc_diff_1 bot_nat_0.not_eq_extremum linorder_le_less_linear not_less_eq zero_neq_one) lemma first1_eq: "first1 xs = length xs - 1" proof - have "tps2 ! j = (\replicate (first1 xs) \ @ [\] @ drop (Suc (first1 xs)) xs\, Suc (Suc (first1 xs)))" (is "_ = (\?zs\, ?i)") using tps2_def jk by simp have "proper_symbols xs" using can bs by fastforce then have *: "proper_symbols ?zs" using proper_symbols_append[of "[\]" "drop (Suc (first1 xs)) xs"] proper_symbols_append by simp have "read tps2 ! j = \?zs\ ?i" using tps2_def jk tapes_at_read'[of j tps2] by simp then have "\?zs\ ?i = \" using read_tps2 by simp then have "?i > length ?zs" using * proper_contents_outofbounds by blast moreover have "length ?zs = length xs" using first1 by simp ultimately have "Suc (first1 xs) \ length xs" by simp moreover have "length xs > 0" using xs by simp ultimately have "first1 xs \ length xs - 1" by simp then show ?thesis using first1(1) by simp qed lemma drop_xs_Nil: "drop (Suc (first1 xs)) xs = []" using first1_eq xs by simp lemma tps2_eq: "tps2 = tps0[j := (\replicate (first1 xs) \ @ [\]\, Suc (Suc (first1 xs)))]" using tps2_def drop_xs_Nil jk by simp definition "tps23 \ tps0 [j := (\replicate (first1 xs) \ @ [\]\, Suc (first1 xs))]" lemma tm23 [transforms_intros]: assumes "ttt = 1" shows "transforms tm23 tps2 ttt tps23" unfolding tm23_def proof (tform tps: tps2_def tps23_def jk time: assms) show "tps23 = tps2[j := tps2 ! j |-| 1]" using tps23_def tps2_eq jk by simp qed definition "tps24 \ tps0 [j := (\replicate (first1 xs) \\, Suc (first1 xs))]" lemma tm24: assumes "ttt = 2" shows "transforms tm24 tps2 ttt tps24" unfolding tm24_def proof (tform tps: tps23_def tps24_def time: assms) show "tps24 = tps23[j := tps23 ! j |:=| \]" proof - have "tps23 ! j |:=| \ = (\replicate (first1 xs) \ @ [\]\, Suc (first1 xs)) |:=| \" using tps23_def jk by simp then have "tps23 ! j |:=| \ = (\replicate (first1 xs) \ @ [\]\, Suc (first1 xs))" using contents_append_update by auto then have "tps23 ! j |:=| \ = (\replicate (first1 xs) \\, Suc (first1 xs))" using contents_append_blanks by (metis replicate_0 replicate_Suc) moreover have "tps24 ! j = (\replicate (first1 xs) \\, Suc (first1 xs))" using tps24_def jk by simp ultimately show ?thesis using tps23_def tps24_def by auto qed qed corollary tm24' [transforms_intros]: assumes "ttt = 2" and "tps' = tps0[j := (\ndecr xs\, Suc (first1 xs))]" shows "transforms tm24 tps2 ttt tps'" proof - have "tps24 = tps0[j := (\ndecr xs\, Suc (first1 xs))]" using tps24_def jk ndecr_def first1_eq xs by simp then show ?thesis using assms tm24 by simp qed end (* context read tps2 ! j = 0 *) context assumes read_tps2': "read tps2 ! j \ \" begin lemma first1_neq: "first1 xs \ length xs - 1" proof (rule ccontr) assume eq: "\ first1 xs \ length xs - 1" have "tps2 ! j = (\replicate (first1 xs) \ @ [\] @ drop (Suc (first1 xs)) xs\, Suc (Suc (first1 xs)))" (is "_ = (\?zs\, ?i)") using tps2_def jk by simp have "length ?zs = length xs" using first1 by simp then have "Suc (Suc (first1 xs)) = Suc (length ?zs)" using xs eq by simp then have *: "\?zs\ ?i = 0" using contents_outofbounds by simp have "read tps2 ! j = \?zs\ ?i" using tps2_def jk tapes_at_read'[of j tps2] by simp then have "\?zs\ ?i \ \" using read_tps2' by simp then show False using * by simp qed lemma tps2: "tps2 = tps0[j := (\ndecr xs\, Suc (Suc (first1 xs)))]" using tps2_def ndecr_def first1_neq xs by simp end (* context read tps2 ! j \ 0 *) lemma tm25 [transforms_intros]: assumes "ttt = (if read tps2 ! j = \ then 4 else 1)" shows "transforms tm25 tps2 ttt tps5" unfolding tm25_def by (tform tps: tps2 tps5_def time: assms) lemma tm5 [transforms_intros]: assumes "ttt = first1 xs + 2 + (if read tps2 ! j = \ then 4 else 1)" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def by (tform time: assms) definition "tps6 \ tps0 [j := (\ndecr xs\, 1)]" lemma tm6: assumes "ttt = first1 xs + 2 + (if read tps2 ! j = \ then 4 else 1) + (tps5 :#: j + 2)" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def proof (tform tps: tps5_def tps6_def jk time: assms) show "clean_tape (tps5 ! j)" proof - have "tps5 ::: j = \ndecr xs\" using tps5_def jk by simp moreover have "bit_symbols (ndecr xs)" using canonical_ndecr can canonical_def by simp ultimately show ?thesis using One_nat_def Suc_1 Suc_le_lessD clean_contents_proper by (metis contents_clean_tape' lessI one_less_numeral_iff semiring_norm(77)) qed qed lemma tm6' [transforms_intros]: assumes "ttt = 2 * first1 xs + 9" shows "transforms tm6 tps0 ttt tps6" proof - let ?ttt = "first1 xs + 2 + (if read tps2 ! j = \ then 4 else 1) + (tps5 :#: j + 2)" have "tps5 :#: j = (if read tps2 ! j = \ then Suc (first1 xs) else Suc (Suc (first1 xs)))" using tps5_def jk by simp then have "?ttt \ ttt" using assms by simp then show ?thesis using tm6 transforms_monotone assms by simp qed end (* context read tps0 ! j \ 0 *) definition "tps7 \ tps0[j := (\ndecr xs\, 1)]" lemma tm7: assumes "ttt = 8 + 2 * length xs" shows "transforms tm7 tps0 ttt tps7" unfolding tm7_def proof (tform tps: tps6_def tps7_def time: assms) show "tps7 = tps0" if "read tps0 ! j = \" using that ndecr_def tps0 tps7_def xs_Nil jk by (simp add: list_update_same_conv) show "2 * first1 xs + 9 + 1 \ ttt" if "read tps0 ! j \ \" proof - have "length xs > 0" using that xs by simp then show ?thesis using first1(1) that assms by simp qed qed end (* context *) end (* locale *) lemma transforms_tm_decrI [transforms_intros]: fixes tps tps' :: "tape list" and n :: nat and k ttt :: nat assumes "j < k" "length tps = k" assumes "tps ! j = (\n\\<^sub>N, 1)" assumes "ttt = 8 + 2 * nlength n" assumes "tps' = tps[j := (\n - 1\\<^sub>N, 1)]" shows "transforms (tm_decr j) tps ttt tps'" proof - let ?xs = "canrepr n" have can: "canonical ?xs" using canonical_canrepr by simp have tps0: "tps ! j = (\?xs\, 1)" using assms by simp have tps': "tps' = tps[j := (\ndecr ?xs\, 1)]" using ndecr assms(5) by (metis canrepr canreprI can canonical_ndecr) interpret loc: turing_machine_decr j . have "transforms loc.tm7 tps ttt tps'" using loc.tm7 loc.tps7_def by (metis assms(1,2,4) can tps' tps0) then show ?thesis using loc.tm7_eq_tm_decr by simp qed subsection \Addition\ text \ In this section we construct a Turing machine that adds two numbers in canonical representation each given on a separate tape and overwrites the second number with the sum. The TM implements the common algorithm with carry starting from the least significant digit. Given two symbol sequences @{term xs} and @{term ys} representing numbers, the next function computes the carry bit that occurs in the $i$-th position. For the least significant position, 0, there is no carry (that is, it is 0); for position $i + 1$ the carry is the sum of the bits of @{term xs} and @{term ys} in position $i$ and the carry for position $i$. The function gives the carry as symbol \textbf{0} or \textbf{1}, except for position 0, where it is the start symbol~$\triangleright$. The start symbol represents the same bit as the symbol~\textbf{0} as defined by @{const todigit}. The reason for this special treatment is that the TM will store the carry on a memorization tape (see~Section~\ref{s:tm-memorizing}), which initially contains the start symbol. \ fun carry :: "symbol list \ symbol list \ nat \ symbol" where "carry xs ys 0 = 1" | "carry xs ys (Suc i) = tosym ((todigit (digit xs i) + todigit (digit ys i) + todigit (carry xs ys i)) div 2)" text \ The next function specifies the $i$-th digit of the sum. \ definition sumdigit :: "symbol list \ symbol list \ nat \ symbol" where "sumdigit xs ys i \ tosym ((todigit (digit xs i) + todigit (digit ys i) + todigit (carry xs ys i)) mod 2)" lemma carry_sumdigit: "todigit (sumdigit xs ys i) + 2 * (todigit (carry xs ys (Suc i))) = todigit (carry xs ys i) + todigit (digit xs i) + todigit (digit ys i)" using sumdigit_def by simp lemma carry_sumdigit_eq_sum: "num xs + num ys = num (map (sumdigit xs ys) [0.. \" proof (induction t) case 0 then show ?case by simp next case (Suc t) then have "todigit (carry xs ys t) \ 1" by simp moreover have "todigit (digit xs t) \ 1" using assms(1) digit_def by auto moreover have "todigit (digit ys t) \ 1" using assms(2) digit_def by auto ultimately show ?case by simp qed lemma num_sumdigit_eq_sum: assumes "length xs \ n" and "length ys \ n" and "symbols_lt 4 xs" and "symbols_lt 4 ys" shows "num xs + num ys = num (map (sumdigit xs ys) [0.. The core of the addition Turing machine is the following command. It scans the symbols on tape $j_1$ and $j_2$ in lockstep until it reaches blanks on both tapes. In every step it adds the symbols on both tapes and the symbol on the last tape, which is a memorization tape storing the carry bit. The sum of these three bits modulo~2 is written to tape $j_2$ and the new carry to the memorization tape. \ definition cmd_plus :: "tapeidx \ tapeidx \ command" where "cmd_plus j1 j2 rs \ (if rs ! j1 = \ \ rs ! j2 = \ then 1 else 0, (map (\j. if j = j1 then (rs ! j, Right) else if j = j2 then (tosym ((todigit (rs ! j1) + todigit (rs ! j2) + todigit (last rs)) mod 2), Right) else if j = length rs - 1 then (tosym ((todigit (rs ! j1) + todigit (rs ! j2) + todigit (last rs)) div 2), Stay) else (rs ! j, Stay)) [0.. j2" and "j1 < k - 1" and "j2 < k - 1" and "j2 > 0" and "length tps = k" and "bit_symbols xs" and "bit_symbols ys" and "tps ! j1 = (\xs\, Suc t)" and "tps ! j2 = (\map (sumdigit xs ys) [0.., Suc t)" and "last tps = \carry xs ys t\" and "rs = read tps" and "tps' = tps [j1 := tps!j1 |+| 1, j2 := tps!j2 |:=| sumdigit xs ys t |+| 1, length tps - 1 := \carry xs ys (Suc t)\]" shows "sem (cmd_plus j1 j2) (0, tps) = (if t < max (length xs) (length ys) then 0 else 1, tps')" proof have "k \ 2" using assms(3,4) by simp have rs1: "rs ! j1 = digit xs t" using assms(2,5,8,11) digit_def read_def contents_def by simp let ?zs = "map (sumdigit xs ys) [0..?zs\ (Suc t) = \" using False contents_def by simp then show ?thesis using digit_def read_def contents_def False assms(3,5,9,11) by simp qed have rs3: "last rs = carry xs ys t" using `k \ 2` assms onesie_read onesie_def read_def read_length tapes_at_read' by (metis (no_types, lifting) diff_less last_conv_nth length_greater_0_conv less_one list.size(3) not_numeral_le_zero) have *: "tosym ((todigit (rs ! j1) + todigit (rs ! j2) + todigit (last rs)) mod 2) = sumdigit xs ys t" using rs1 rs2 rs3 sumdigit_def by simp have "\ (digit xs t = 0 \ digit ys t = 0)" if "t < max (length xs) (length ys)" using assms(6,7) digit_def that by auto then have 4: "\ (rs ! j1 = 0 \ rs ! j2 = 0)" if "t < max (length xs) (length ys)" using rs1 rs2 that by simp then have fst1: "fst (sem (cmd_plus j1 j2) (0, tps)) = fst (0, tps')" if "t < max (length xs) (length ys)" using that cmd_plus_def assms(11) by (smt (verit, ccfv_threshold) fst_conv prod.sel(2) sem) have "digit xs t = 0 \ digit ys t = 0" if "t \ max (length xs) (length ys)" using that digit_def by simp then have 5: "rs ! j1 = \ \ rs ! j2 = \" if "t \ max (length xs) (length ys)" using rs1 rs2 that by simp then have "fst (sem (cmd_plus j1 j2) (0, tps)) = fst (1, tps')" if "t \ max (length xs) (length ys)" using that cmd_plus_def assms(11) by (smt (verit, ccfv_threshold) fst_conv prod.sel(2) sem) then show "fst (sem (cmd_plus j1 j2) (0, tps)) = fst (if t < max (length xs) (length ys) then 0 else 1, tps')" using fst1 by (simp add: not_less) show "snd (sem (cmd_plus j1 j2) (0, tps)) = snd (if t < max (length xs) (length ys) then 0 else 1, tps')" proof (rule snd_semI) show "proper_command k (cmd_plus j1 j2)" using cmd_plus_def by simp show "length tps = k" using assms(5) . show "length tps' = k" using assms(5,12) by simp have len: "length (read tps) = k" by (simp add: assms read_length) show "act (cmd_plus j1 j2 (read tps) [!] j) (tps ! j) = tps' ! j" if "j < k" for j proof - have j: "j < length tps" using len that assms(5) by simp consider "j = j1" | "j \ j1 \ j = j2" | "j \ j1 \ j \ j2 \ j = length rs - 1" | "j \ j1 \ j \ j2 \ j \ length rs - 1" by auto then show ?thesis proof (cases) case 1 then have "cmd_plus j1 j2 (read tps) [!] j = (read tps ! j, Right)" using that len cmd_plus_def by simp then have "act (cmd_plus j1 j2 (read tps) [!] j) (tps ! j) = tps ! j |+| 1" using act_Right[OF j] by simp moreover have "tps' ! j = tps ! j |+| 1" using assms(1,2,5,12) that 1 by simp ultimately show ?thesis by simp next case 2 then have "cmd_plus j1 j2 (read tps) [!] j = (tosym ((todigit (rs ! j1) + todigit (rs ! j2) + todigit (last rs)) mod 2), Right)" using that len cmd_plus_def assms(11) by simp then have "cmd_plus j1 j2 (read tps) [!] j = (sumdigit xs ys t, Right)" using * by simp moreover have "tps' ! j2 = tps!j2 |:=| sumdigit xs ys t |+| 1" using assms(3,5,12) by simp ultimately show ?thesis using act_Right' 2 by simp next case 3 then have "cmd_plus j1 j2 (read tps) [!] j = (tosym ((todigit (rs ! j1) + todigit (rs ! j2) + todigit (last rs)) div 2), Stay)" using that len cmd_plus_def assms(11) by simp then have "cmd_plus j1 j2 (read tps) [!] j = (carry xs ys (Suc t), Stay)" using rs1 rs2 rs3 by simp moreover have "tps' ! (length tps - 1) = \carry xs ys (Suc t)\" using 3 assms(5,11,12) len that by simp ultimately show ?thesis using 3 act_onesie assms(3,5,10,11) len by (metis add_diff_inverse_nat last_length less_nat_zero_code nat_diff_split_asm plus_1_eq_Suc) next case 4 then have "cmd_plus j1 j2 (read tps) [!] j = (read tps ! j, Stay)" using that len cmd_plus_def assms(11) by simp then have "act (cmd_plus j1 j2 (read tps) [!] j) (tps ! j) = tps ! j" using act_Stay[OF j] by simp moreover have "tps' ! j = tps ! j" using that 4 len assms(5,11,12) by simp ultimately show ?thesis by simp qed qed qed qed lemma contents_map_append_drop: "\map f [0..(Suc t := f t) = \map f [0.." proof (cases "t < length zs") case lt: True then have t_lt: "t < length (map f [0.. 0 \ x < Suc t" | "x = Suc t" | "x > Suc t \ x \ length zs" | "x > Suc t \ x > length zs" by linarith then show "(\map f [0..(Suc t := f t)) x = \map f [0.. x" (is "?lhs x = ?rhs x") proof (cases) case 1 then show ?thesis using contents_def by simp next case 2 then have "?lhs x = (map f [0..map f [0.. x" using contents_def by simp then have "?lhs x = (map f [0.. length (map f [0.. 0" using 4 by simp ultimately have "?rhs x = (map f [0.. Suc t" using 4 by auto ultimately have "?rhs x = drop (Suc t) zs ! (x - 1 - Suc t)" by (metis diff_zero leD length_map length_upt nth_append) then show ?thesis using left 4 by (metis Cons_nth_drop_Suc Suc_diff_Suc diff_Suc_eq_diff_pred lt nth_Cons_Suc) next case 5 then show ?thesis using lt contents_def by auto qed qed next case False moreover have "\map f [0..(Suc t := f t) = \map f [0.." proof fix x show "(\map f [0..(Suc t := f t)) x = \map f [0.. x" proof (cases "x < Suc t") case True then show ?thesis using contents_def by (smt (verit, del_insts) diff_Suc_1 diff_zero fun_upd_apply length_map length_upt less_Suc_eq_0_disj less_Suc_eq_le less_imp_le_nat nat_neq_iff nth_map_upt) next case ge: False show ?thesis proof (cases "x = Suc t") case True then show ?thesis using contents_def by (metis One_nat_def add_Suc diff_Suc_1 diff_zero fun_upd_same ge le_eq_less_or_eq length_map length_upt lessI less_Suc_eq_0_disj nth_map_upt plus_1_eq_Suc) next case False then have "x > Suc t" using ge by simp then show ?thesis using contents_def by simp qed qed qed ultimately show ?thesis by simp qed corollary sem_cmd_plus': assumes "j1 \ j2" and "j1 < k - 1" and "j2 < k - 1" and "j2 > 0" and "length tps = k" and "bit_symbols xs" and "bit_symbols ys" and "tps ! j1 = (\xs\, Suc t)" and "tps ! j2 = (\map (sumdigit xs ys) [0.., Suc t)" and "last tps = \carry xs ys t\" and "tps' = tps [j1 := (\xs\, Suc (Suc t)), j2 := (\map (sumdigit xs ys) [0.., Suc (Suc t)), length tps - 1 := \carry xs ys (Suc t)\]" shows "sem (cmd_plus j1 j2) (0, tps) = (if Suc t \ max (length xs) (length ys) then 0 else 1, tps')" proof - have "tps ! j1 |+| 1 = (\xs\, Suc (Suc t))" using assms(8) by simp moreover have "tps ! j2 |:=| sumdigit xs ys t |+| 1 = (\map (sumdigit xs ys) [0.., Suc (Suc t))" using contents_map_append_drop assms(9) by simp ultimately show ?thesis using sem_cmd_plus[OF assms(1-10)] assms(11) by auto qed text \ The next Turing machine comprises just the command @{const cmd_plus}. It overwrites tape $j_2$ with the sum of the numbers on tape $j_1$ and $j_2$. The carry bit is maintained on the last tape. \ definition tm_plus :: "tapeidx \ tapeidx \ machine" where "tm_plus j1 j2 \ [cmd_plus j1 j2]" lemma tm_plus_tm: assumes "j2 > 0" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_plus j1 j2)" unfolding tm_plus_def using assms(1-3) cmd_plus_def turing_machine_def by auto lemma tm_plus_immobile: fixes k :: nat assumes "j1 < k" and "j2 < k" shows "immobile (tm_plus j1 j2) k (Suc k)" proof - let ?M = "tm_plus j1 j2" { fix q :: nat and rs :: "symbol list" assume q: "q < length ?M" assume rs: "length rs = Suc k" then have len: "length rs - 1 = k" by simp have neq: "k \ j1" "k \ j2" using assms by simp_all have "?M ! q = cmd_plus j1 j2" using tm_plus_def q by simp moreover have "(cmd_plus j1 j2) rs [!] k = (tosym ((todigit (rs ! j1) + todigit (rs ! j2) + todigit (last rs)) div 2), Stay)" using cmd_plus_def rs len neq by fastforce ultimately have "(cmd_plus j1 j2) rs [~] k = Stay" by simp } then show ?thesis by (simp add: immobile_def tm_plus_def) qed lemma execute_tm_plus: assumes "j1 \ j2" and "j1 < k - 1" and "j2 < k - 1" and "j2 > 0" and "length tps = k" and "bit_symbols xs" and "bit_symbols ys" and "t \ Suc (max (length xs) (length ys))" and "tps ! j1 = (\xs\, 1)" and "tps ! j2 = (\ys\, 1)" and "last tps = \\\" shows "execute (tm_plus j1 j2) (0, tps) t = (if t \ max (length xs) (length ys) then 0 else 1, tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length tps - 1 := \carry xs ys t\])" using assms(8) proof (induction t) case 0 have "carry xs ys 0 = 1" by simp moreover have "map (sumdigit xs ys) [0..<0] @ drop 0 ys = ys" by simp ultimately have "tps = tps [j1 := (\xs\, Suc 0), j2 := (\map (sumdigit xs ys) [0..<0] @ drop 0 ys\, Suc 0), length tps - 1 := \carry xs ys 0\]" using assms by (metis One_nat_def add_diff_inverse_nat last_length less_nat_zero_code list_update_id nat_diff_split_asm plus_1_eq_Suc) then show ?case by simp next case (Suc t) let ?M = "tm_plus j1 j2" have "execute ?M (0, tps) (Suc t) = exe ?M (execute ?M (0, tps) t)" (is "_ = exe ?M ?cfg") by simp also have "... = sem (cmd_plus j1 j2) ?cfg" using Suc tm_plus_def exe_lt_length by simp also have "... = (if Suc t \ max (length xs) (length ys) then 0 else 1, tps [j1 := (\xs\, Suc (Suc t)), j2 := (\map (sumdigit xs ys) [0.., Suc (Suc t)), length tps - 1 := \carry xs ys (Suc t)\])" proof - let ?tps = "tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length tps - 1 := \carry xs ys t\]" let ?tps' = "?tps [j1 := (\xs\, Suc (Suc t)), j2 := (\map (sumdigit xs ys) [0.., Suc (Suc t)), length tps - 1 := \carry xs ys (Suc t)\]" have cfg: "?cfg = (0, ?tps)" using Suc by simp have tps_k: "length ?tps = k" using assms(2,3,5) by simp have tps_j1: "?tps ! j1 = (\xs\, Suc t)" using assms(1-3,5) by simp have tps_j2: "?tps ! j2 = (\map (sumdigit xs ys) [0.., Suc t)" using assms(1-3,5) by simp have tps_last: "last ?tps = \carry xs ys t\" using assms by (metis One_nat_def carry.simps(1) diff_Suc_1 last_list_update length_list_update list_update_nonempty prod.sel(2) tps_j1) then have "sem (cmd_plus j1 j2) (0, ?tps) = (if Suc t \ max (length xs) (length ys) then 0 else 1, ?tps')" using sem_cmd_plus'[OF assms(1-4) tps_k assms(6,7) tps_j1 tps_j2 tps_last] assms(1-3) by (smt (verit, best) Suc.prems Suc_lessD assms(5) tps_k) then have "sem (cmd_plus j1 j2) ?cfg = (if Suc t \ max (length xs) (length ys) then 0 else 1, ?tps')" using cfg by simp moreover have "?tps' = tps [j1 := (\xs\, Suc (Suc t)), j2 := (\map (sumdigit xs ys) [0.., Suc (Suc t)), length tps - 1 := \carry xs ys (Suc t)\]" - using assms by (smt (z3) list_update_overwrite list_update_swap) + using assms by (smt (verit) list_update_overwrite list_update_swap) ultimately show ?thesis by simp qed finally show ?case by simp qed lemma tm_plus_bounded_write: assumes "j1 < k - 1" shows "bounded_write (tm_plus j1 j2) (k - 1) 4" using assms cmd_plus_def tm_plus_def bounded_write_def by simp lemma carry_max_length: assumes "bit_symbols xs" and "bit_symbols ys" shows "carry xs ys (Suc (max (length xs) (length ys))) = \" proof - let ?t = "max (length xs) (length ys)" have "carry xs ys (Suc ?t) = tosym ((todigit (digit xs ?t) + todigit (digit ys ?t) + todigit (carry xs ys ?t)) div 2)" by simp then have "carry xs ys (Suc ?t) = tosym (todigit (carry xs ys ?t) div 2)" using digit_def by simp moreover have "carry xs ys ?t \ \" using carry_le assms by fastforce ultimately show ?thesis by simp qed corollary execute_tm_plus_halt: assumes "j1 \ j2" and "j1 < k - 1" and "j2 < k - 1" and "j2 > 0" and "length tps = k" and "bit_symbols xs" and "bit_symbols ys" and "t = Suc (max (length xs) (length ys))" and "tps ! j1 = (\xs\, 1)" and "tps ! j2 = (\ys\, 1)" and "last tps = \\\" shows "execute (tm_plus j1 j2) (0, tps) t = (1, tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length tps - 1 := \\\])" proof - have "execute (tm_plus j1 j2) (0, tps) t = (1, tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length tps - 1 := \carry xs ys t\])" using assms(8) execute_tm_plus[OF assms(1-7) _ assms(9-11)] Suc_leI Suc_n_not_le_n lessI by presburger then have "execute (tm_plus j1 j2) (0, tps) t = (1, tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length tps - 1 := \carry xs ys t\])" using assms(8) by simp then show "execute (tm_plus j1 j2) (0, tps) t = (1, tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length tps - 1 := \\\])" using assms(8) carry_max_length[OF assms(6,7)] by metis qed lemma transforms_tm_plusI: assumes "j1 \ j2" and "j1 < k - 1" and "j2 < k - 1" and "j2 > 0" and "length tps = k" and "bit_symbols xs" and "bit_symbols ys" and "t = Suc (max (length xs) (length ys))" and "tps ! j1 = (\xs\, 1)" and "tps ! j2 = (\ys\, 1)" and "last tps = \\\" and "tps' = tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length tps - 1 := \\\]" shows "transforms (tm_plus j1 j2) tps t tps'" using assms execute_tm_plus_halt[OF assms(1-11)] tm_plus_def transforms_def transits_def by auto text \ The next Turing machine removes the memorization tape from @{const tm_plus}. \ definition tm_plus' :: "tapeidx \ tapeidx \ machine" where "tm_plus' j1 j2 \ cartesian (tm_plus j1 j2) 4" lemma tm_plus'_tm: assumes "j2 > 0" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_plus' j1 j2)" unfolding tm_plus'_def using assms cartesian_tm tm_plus_tm by simp lemma transforms_tm_plus'I [transforms_intros]: fixes k t :: nat and j1 j2 :: tapeidx and tps tps' :: "tape list" and xs zs :: "symbol list" assumes "j1 \ j2" and "j1 < k" and "j2 < k" and "j2 > 0" and "length tps = k" and "bit_symbols xs" and "bit_symbols ys" and "t = Suc (max (length xs) (length ys))" and "tps ! j1 = (\xs\, 1)" and "tps ! j2 = (\ys\, 1)" and "tps' = tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t)]" shows "transforms (tm_plus' j1 j2) tps t tps'" proof - let ?tps = "tps @ [\\\]" let ?tps' = "?tps [j1 := (\xs\, Suc t), j2 := (\map (sumdigit xs ys) [0.., Suc t), length ?tps - 1 := \\\]" let ?M = "tm_plus j1 j2" have 1: "length ?tps = Suc k" using assms(5) by simp have 2: "?tps ! j1 = (\xs\, 1)" by (simp add: assms(9) assms(2) assms(5) nth_append) have 3: "?tps ! j2 = (\ys\, 1)" by (simp add: assms(10) assms(3) assms(5) nth_append) have 4: "last ?tps = \\\" by simp have 5: "k \ 2" using assms(3,4) by simp have "transforms (tm_plus j1 j2) ?tps t ?tps'" using transforms_tm_plusI[OF assms(1) _ _ assms(4) 1 assms(6,7,8) 2 3 4, of ?tps'] assms(2,3) by simp moreover have "?tps' = tps' @ [\\\]" using assms by (simp add: list_update_append) ultimately have "transforms (tm_plus j1 j2) (tps @ [\\\]) t (tps' @ [\\\])" by simp moreover have "turing_machine (Suc k) 4 ?M" using tm_plus_tm assms by simp moreover have "immobile ?M k (Suc k)" using tm_plus_immobile assms by simp moreover have "bounded_write (tm_plus j1 j2) k 4" using tm_plus_bounded_write[of j1 "Suc k"] assms(2) by simp ultimately have "transforms (cartesian (tm_plus j1 j2) 4) tps t tps'" using cartesian_transforms_onesie[where ?M="?M" and ?b=4] assms(5) 5 by simp then show ?thesis using tm_plus'_def by simp qed text \ The next Turing machine is the one we actually use to add two numbers. After computing the sum by running @{const tm_plus'}, it removes trailing zeros and performs a carriage return on the tapes $j_1$ and $j_2$. \ definition tm_add :: "tapeidx \ tapeidx \ machine" where "tm_add j1 j2 \ tm_plus' j1 j2 ;; tm_lconst_until j2 j2 {h. h \ \ \ h \ \} \ ;; tm_cr j1 ;; tm_cr j2" lemma tm_add_tm: assumes "j2 > 0" and "k \ 2" and "G \ 4" and "j2 < k" shows "turing_machine k G (tm_add j1 j2)" unfolding tm_add_def using tm_plus'_tm tm_lconst_until_tm tm_cr_tm assms by simp locale turing_machine_add = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_plus' j1 j2" definition "tm2 \ tm1 ;; tm_lconst_until j2 j2 {h. h \ \ \ h \ \} \" definition "tm3 \ tm2 ;; tm_cr j1" definition "tm4 \ tm3 ;; tm_cr j2" lemma tm4_eq_tm_add: "tm4 = tm_add j1 j2" using tm4_def tm3_def tm2_def tm1_def tm_add_def by simp context fixes x y k :: nat and tps0 :: "tape list" assumes jk: "j1 \ j2" "j1 < k" "j2 < k" "j2 > 0" "k = length tps0" assumes tps0: "tps0 ! j1 = (\canrepr x\, 1)" "tps0 ! j2 = (\canrepr y\, 1)" begin abbreviation "xs \ canrepr x" abbreviation "ys \ canrepr y" lemma xs: "bit_symbols xs" using bit_symbols_canrepr by simp lemma ys: "bit_symbols ys" using bit_symbols_canrepr by simp abbreviation "n \ Suc (max (length xs) (length ys))" abbreviation "m \ length (canrepr (num xs + num ys))" definition "tps1 \ tps0 [j1 := (\xs\, Suc n), j2 := (\map (sumdigit xs ys) [0.., Suc n)]" lemma tm1 [transforms_intros]: assumes "ttt = n" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: jk xs ys tps0 time: assms) show "tps1 = tps0 [j1 := (\xs\, Suc ttt), j2 := (\map (sumdigit xs ys) [0.., Suc ttt)]" using tps1_def assms by simp qed definition "tps2 \ tps0 [j1 := (\xs\, Suc n), j2 := (\canrepr (num xs + num ys)\, m)]" lemma contents_canlen: assumes "bit_symbols zs" shows "\zs\ (canlen zs) \ {h. h \ \ \ \ < h}" using assms contents_def canlen_le_length canlen_one by auto lemma tm2 [transforms_intros]: assumes "ttt = n + Suc (Suc n - canlen (map (sumdigit xs ys) [0..?zs\, Suc n) {h. h \ \ \ \ < h} ?ln" proof (rule lneighI) have "\?zs\ (canlen ?zs) \ {h. h \ \ \ \ < h}" using contents_canlen[OF `bit_symbols ?zs`] by simp moreover have "Suc n - ?ln = canlen ?zs" by (metis One_nat_def diff_Suc_1 diff_Suc_Suc diff_diff_cancel le_imp_less_Suc length_map length_upt less_imp_le_nat canlen_le_length) ultimately have "\?zs\ (Suc n - ?ln) \ {h. h \ \ \ \ < h}" by simp then show "fst (\?zs\, Suc n) (snd (\?zs\, Suc n) - ?ln) \ {h. h \ \ \ \ < h}" by simp have "\?zs\ (Suc n - n') \ {\, \}" if "n' < ?ln" for n' proof (cases "Suc n - n' \ n") case True moreover have 1: "Suc n - n' > 0" using that by simp ultimately have "\?zs\ (Suc n - n') = ?zs ! (Suc n - n' - 1)" using contents_def by simp moreover have "Suc n - n' - 1 < length ?zs" using that True by simp moreover have "Suc n - n' - 1 \ canlen ?zs" using that by simp ultimately show ?thesis using canlen_at_ge[of ?zs] by simp next case False then show ?thesis by simp qed then have "\?zs\ (Suc n - n') \ {h. h \ \ \ \ < h}" if "n' < ?ln" for n' using that by fastforce then show "fst (\?zs\, Suc n) (snd (\?zs\, Suc n) - n') \ {h. h \ \ \ \ < h}" if "n' < ?ln" for n' using that by simp qed then show "lneigh (tps1 ! j2) {h. h \ \ \ h \ \} ?ln" using assms tps1_def jk by simp show "Suc n - canlen (map (sumdigit xs ys) [0.. tps1 :#: j2" "Suc n - canlen (map (sumdigit xs ys) [0.. tps1 :#: j2" using assms tps1_def jk by simp_all have num_zs: "num ?zs = num xs + num ys" using assms num_sumdigit_eq_sum'' xs ys by simp then have canrepr: "canrepr (num xs + num ys) = take (canlen ?zs) ?zs" using canrepr_take_canlen `bit_symbols ?zs` by blast have len_canrepr: "length (canrepr (num xs + num ys)) = canlen ?zs" using num_zs length_canrepr_canlen sumdigit_bit_symbols by blast have "lconstplant (\?zs\, Suc n) \ ?ln = (\canrepr (num xs + num ys)\, m)" (is "lconstplant ?tp \ ?ln = _") proof - have "(if Suc n - ?ln < i \ i \ Suc n then \ else \?zs\ i) = \take (canlen ?zs) ?zs\ i" (is "?lhs = ?rhs") for i proof - consider "i = 0" | "i > 0 \ i \ canlen ?zs" | "i > canlen ?zs \ i \ Suc n" | "i > canlen ?zs \ i > Suc n" by linarith then show ?thesis proof (cases) case 1 then show ?thesis by simp next case 2 then have "i \ Suc n - ?ln" using canlen_le_length by (metis diff_diff_cancel diff_zero le_imp_less_Suc length_map length_upt less_imp_le_nat) then have lhs: "?lhs = \?zs\ i" by simp have "take (canlen ?zs) ?zs ! (i - 1) = ?zs ! (i - 1)" using 2 by (metis Suc_diff_1 Suc_less_eq le_imp_less_Suc nth_take) then have "?rhs = \?zs\ i" using 2 contents_inbounds len_canrepr local.canrepr not_le canlen_le_length by (metis add_diff_inverse_nat add_leE) then show ?thesis using lhs by simp next case 3 then have "Suc n - ?ln < i \ i \ Suc n" by (metis diff_diff_cancel less_imp_le_nat less_le_trans) then have "?lhs = 0" by simp moreover have "?rhs = 0" using 3 contents_outofbounds len_canrepr canrepr by metis ultimately show ?thesis by simp next case 4 then have "?lhs = 0" by simp moreover have "?rhs = 0" using 4 contents_outofbounds len_canrepr canrepr by metis ultimately show ?thesis by simp qed qed then have "(\i. if Suc n - ?ln < i \ i \ Suc n then \ else \?zs\ i) = \canrepr (num xs + num ys)\" using canrepr by simp moreover have "fst ?tp = \?zs\" by simp ultimately have "(\i. if Suc n - ?ln < i \ i \ Suc n then 0 else fst ?tp i) = \canrepr (num xs + num ys)\" by metis moreover have "Suc n - ?ln = m" using len_canrepr by (metis add_diff_inverse_nat diff_add_inverse2 diff_is_0_eq diff_zero le_imp_less_Suc length_map length_upt less_imp_le_nat less_numeral_extra(3) canlen_le_length zero_less_diff) ultimately show ?thesis using lconstplant[of ?tp 0 ?ln] by simp qed then show "tps2 = tps1 [j2 := tps1 ! j2 |-| ?ln, j2 := lconstplant (tps1 ! j2) 0 ?ln]" using tps2_def tps1_def jk by simp show "ttt = n + Suc ?ln" using assms by simp qed definition "tps3 \ tps0 [j1 := (\xs\, 1), j2 := (\canrepr (num xs + num ys)\, m)]" lemma tm3 [transforms_intros]: assumes "ttt = n + Suc (Suc n - canlen (map (sumdigit xs ys) [0.. tps0 [j1 := (\xs\, 1), j2 := (\canrepr (num xs + num ys)\, 1)]" lemma tm4: assumes "ttt = n + Suc (Suc n - canlen (map (sumdigit xs ys) [0.. n" by (metis canlen_le_length) have "n + Suc (Suc n - canlen ?zs) + Suc n + 2 + m + 2 = n + Suc (Suc n - m) + Suc n + 2 + m + 2" using 1 by simp also have "... = n + Suc (Suc n - m) + Suc n + 4 + m" by simp also have "... = n + Suc (Suc n) - m + Suc n + 4 + m" using `m \ n` by simp also have "... = n + Suc (Suc n) + Suc n + 4" using `m \ n` by simp also have "... = 3 * n + 7" by simp also have "... = ttt" using assms by simp finally have "n + Suc (Suc n - canlen ?zs) + Suc n + 2 + m + 2 = ttt" . then show ?thesis using tm4 by simp qed definition "tps4' \ tps0 [j2 := (\x + y\\<^sub>N, 1)]" lemma tm4'': assumes "ttt = 3 * max (nlength x) (nlength y) + 10" shows "transforms tm4 tps0 ttt tps4'" proof - have "canrepr (num xs + num ys) = canrepr (x + y)" by (simp add: canrepr) then show ?thesis using assms tps0(1) tps4'_def tps4_def tm4' by (metis list_update_id) qed end (* context *) end (* locale *) lemma transforms_tm_addI [transforms_intros]: fixes j1 j2 :: tapeidx fixes x y k ttt :: nat and tps tps' :: "tape list" assumes "j1 \ j2" "j1 < k" "j2 < k" "j2 > 0" "k = length tps" assumes "tps ! j1 = (\canrepr x\, 1)" "tps ! j2 = (\canrepr y\, 1)" assumes "ttt = 3 * max (nlength x) (nlength y) + 10" assumes "tps' = tps [j2 := (\x + y\\<^sub>N, 1)]" shows "transforms (tm_add j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_add j1 j2 . show ?thesis using loc.tm4_eq_tm_add loc.tps4'_def loc.tm4'' assms by simp qed subsection \Multiplication\ text \ In this section we construct a Turing machine that multiplies two numbers, each on its own tape, and writes the result to another tape. It employs the common algorithm for multiplication, which for binary numbers requires only doubling a number and adding two numbers. For the latter we already have a TM; for the former we are going to construct one. \ subsubsection \The common algorithm\ text \ For two numbers given as symbol sequences @{term xs} and @{term ys}, the common algorithm maintains an intermediate result, initialized with 0, and scans @{term xs} starting from the most significant digit. In each step the intermediate result is multiplied by two, and if the current digit of @{term xs} is @{text \}, the value of @{term ys} is added to the intermediate result. \ fun prod :: "symbol list \ symbol list \ nat \ nat" where "prod xs ys 0 = 0" | "prod xs ys (Suc i) = 2 * prod xs ys i + (if xs ! (length xs - 1 - i) = 3 then num ys else 0)" text \ After $i$ steps of the algorithm, the intermediate result is the product of @{term ys} and the $i$ most significant bits of @{term xs}. \ lemma prod: assumes "i \ length xs" shows "prod xs ys i = num (drop (length xs - i) xs) * num ys" using assms proof (induction i) case 0 then show ?case using num_def by simp next case (Suc i) then have "i < length xs" by simp then have "drop (length xs - Suc i) xs = (xs ! (length xs - 1 - i)) # drop (length xs - i) xs" by (metis Cons_nth_drop_Suc Suc_diff_Suc diff_Suc_eq_diff_pred diff_Suc_less gr_implies_not0 length_greater_0_conv list.size(3)) then show ?case using num_Cons Suc by simp qed text \ After @{term "length xs"} steps, the intermediate result is the final result: \ corollary prod_eq_prod: "prod xs ys (length xs) = num xs * num ys" using prod by simp definition prod' :: "nat \ nat \ nat \ nat" where "prod' x y i \ prod (canrepr x) (canrepr y) i" lemma prod': "prod' x y (nlength x) = x * y" using prod_eq_prod prod'_def by (simp add: canrepr) subsubsection \Multiplying by two\ text \ Since we represent numbers with the least significant bit at the left, a multiplication by two is a right shift with a \textbf{0} inserted as the least significant digit. The next command implements the right shift. It scans the tape $j$ and memorizes the current symbol on the last tape. It only writes the symbols \textbf{0} and \textbf{1}. \ definition cmd_double :: "tapeidx \ command" where "cmd_double j rs \ (if rs ! j = \ then 1 else 0, (map (\i. if i = j then if last rs = \ \ rs ! j = \ then (rs ! i, Right) else (tosym (todigit (last rs)), Right) else if i = length rs - 1 then (tosym (todigit (rs ! j)), Stay) else (rs ! i, Stay)) [0.. 2" and "G \ 4" and "j > 0" and "j < k - 1" shows "turing_command k 1 G (cmd_double j)" proof show "\gs. length gs = k \ length ([!!] cmd_double j gs) = length gs" using cmd_double_def by simp show "\gs. length gs = k \ 0 < k \ cmd_double j gs [.] 0 = gs ! 0" using assms cmd_double_def by simp show "cmd_double j gs [.] j' < G" if "length gs = k" "\i. i < length gs \ gs ! i < G" "j' < length gs" for j' gs proof - consider "j' = j" | "j' = k - 1" | "j' \ j \ j' \ k - 1" by auto then show ?thesis proof (cases) case 1 then have "cmd_double j gs [!] j' = (if last gs = \ \ gs ! j = \ then (gs ! j, Right) else (tosym (todigit (last gs)), Right))" using cmd_double_def assms(1,4) that(1) by simp then have "cmd_double j gs [.] j' = (if last gs = \ \ gs ! j = \ then gs ! j else tosym (todigit (last gs)))" by simp then show ?thesis using that assms by simp next case 2 then have "cmd_double j gs [!] j' = (tosym (todigit (gs ! j)), Stay)" using cmd_double_def assms(1,4) that(1) by simp then show ?thesis using assms by simp next case 3 then show ?thesis using cmd_double_def assms that by simp qed qed show "\gs. length gs = k \ [*] (cmd_double j gs) \ 1" using assms cmd_double_def by simp qed lemma sem_cmd_double_0: assumes "j < k" and "bit_symbols xs" and "i \ length xs" and "i > 0" and "length tps = Suc k" and "tps ! j = (\xs\, i)" and "tps ! k = \z\" and "tps' = tps [j := tps ! j |:=| tosym (todigit z) |+| 1, k := \xs ! (i - 1)\]" shows "sem (cmd_double j) (0, tps) = (0, tps')" proof (rule semI) show "proper_command (Suc k) (cmd_double j)" using cmd_double_def by simp show "length tps = Suc k" using assms(5) . show "length tps' = Suc k" using assms(5,8) by simp show "fst (cmd_double j (read tps)) = 0" using assms contents_def cmd_double_def tapes_at_read'[of j tps] by (smt (verit, del_insts) One_nat_def Suc_le_lessD Suc_le_mono Suc_pred fst_conv less_imp_le_nat snd_conv zero_neq_numeral) show "act (cmd_double j (read tps) [!] j') (tps ! j') = tps' ! j'" if "j' < Suc k" for j' proof - define rs where "rs = read tps" then have rsj: "rs ! j = xs ! (i - 1)" using assms tapes_at_read' contents_inbounds by (metis fst_conv le_imp_less_Suc less_imp_le_nat snd_conv) then have rs23: "rs ! j = \ \ rs ! j = \" using assms by simp have lenrs: "length rs = Suc k" by (simp add: rs_def assms(5) read_length) consider "j' = j" | "j' = k" | "j' \ j \ j' \ k" by auto then show ?thesis proof (cases) case 1 then have "j' < length rs" using lenrs that by simp then have "cmd_double j rs [!] j' = (if last rs = \ \ rs ! j = \ then (rs ! j, Right) else (tosym (todigit (last rs)), Right))" using cmd_double_def that 1 by simp then have "cmd_double j rs [!] j' = (tosym (todigit (last rs)), Right)" using rs23 lenrs assms by auto moreover have "last rs = z" using lenrs assms(5,7) rs_def onesie_read[of z] tapes_at_read'[of _ tps] by (metis diff_Suc_1 last_conv_nth length_0_conv lessI old.nat.distinct(2)) ultimately show ?thesis using act_Right' rs_def 1 assms(1,5,8) by simp next case 2 then have "j' = length rs - 1" "j' \ j" "j' < length rs" using lenrs that assms(1) by simp_all then have "(cmd_double j rs) [!] j' = (tosym (todigit (rs ! j)), Stay)" using cmd_double_def by simp then have "(cmd_double j rs) [!] j' = (xs ! (i - 1), Stay)" using rsj rs23 by auto then show ?thesis using act_onesie rs_def 2 assms that by simp next case 3 then have "j' \ length rs - 1" "j' \ j" "j' < length rs" using lenrs that by simp_all then have "(cmd_double j rs) [!] j' = (rs ! j', Stay)" using cmd_double_def by simp then show ?thesis using act_Stay rs_def assms that 3 by simp qed qed qed lemma sem_cmd_double_1: assumes "j < k" and "bit_symbols xs" and "i > length xs" and "length tps = Suc k" and "tps ! j = (\xs\, i)" and "tps ! k = \z\" and "tps' = tps [j := tps ! j |:=| (if z = \ then \ else tosym (todigit z)) |+| 1, k := \\\]" shows "sem (cmd_double j) (0, tps) = (1, tps')" proof (rule semI) show "proper_command (Suc k) (cmd_double j)" using cmd_double_def by simp show "length tps = Suc k" using assms(4) . show "length tps' = Suc k" using assms(4,7) by simp show "fst (cmd_double j (read tps)) = 1" using assms contents_def cmd_double_def tapes_at_read'[of j tps] by simp have "j < length tps" using assms by simp show "act (cmd_double j (read tps) [!] j') (tps ! j') = tps' ! j'" if "j' < Suc k" for j' proof - define rs where "rs = read tps" then have rsj: "rs ! j = \" using tapes_at_read'[OF `j < length tps`] assms(1,3,4,5) by simp have lenrs: "length rs = Suc k" by (simp add: rs_def assms(4) read_length) consider "j' = j" | "j' = k" | "j' \ j \ j' \ k" by auto then show ?thesis proof (cases) case 1 then have "j' < length rs" using lenrs that by simp then have "cmd_double j rs [!] j' = (if last rs = \ \ rs ! j = \ then (rs ! j, Right) else (tosym (todigit (last rs)), Right))" using cmd_double_def that 1 by simp moreover have "last rs = z" using assms onesie_read rs_def tapes_at_read' by (metis diff_Suc_1 last_conv_nth length_0_conv lenrs lessI nat.simps(3)) ultimately have "cmd_double j rs [!] j' = (if z = \ then (\, Right) else (tosym (todigit z), Right))" using rsj 1 by simp then show ?thesis using act_Right' rs_def 1 assms(1,4,7) by simp next case 2 then have "j' = length rs - 1" "j' \ j" "j' < length rs" using lenrs that assms(1) by simp_all then have "(cmd_double j rs) [!] j' = (tosym (todigit (rs ! j)), Stay)" using cmd_double_def by simp then have "(cmd_double j rs) [!] j' = (2, Stay)" using rsj by auto then show ?thesis using act_onesie rs_def 2 assms that by simp next case 3 then have "j' \ length rs - 1" "j' \ j" "j' < length rs" using lenrs that by simp_all then have "(cmd_double j rs) [!] j' = (rs ! j', Stay)" using cmd_double_def by simp then show ?thesis using act_Stay rs_def assms that 3 by simp qed qed qed text \ The next Turing machine consists just of the command @{const cmd_double}. \ definition tm_double :: "tapeidx \ machine" where "tm_double j \ [cmd_double j]" lemma tm_double_tm: assumes "k \ 2" and "G \ 4" and "j > 0" and "j < k - 1" shows "turing_machine k G (tm_double j)" using assms tm_double_def turing_command_double by auto lemma execute_tm_double_0: assumes "j < k" and "bit_symbols xs" and "length xs > 0" and "length tps = Suc k" and "tps ! j = (\xs\, 1)" and "tps ! k = \\\" and "t \ 1" and "t \ length xs" shows "execute (tm_double j) (0, tps) t = (0, tps [j := (\\ # take (t - 1) xs @ drop t xs\, Suc t), k := \xs ! (t - 1)\])" using assms(7,8) proof (induction t rule: nat_induct_at_least) case base have "execute (tm_double j) (0, tps) 1 = exe (tm_double j) (execute (tm_double j) (0, tps) 0)" by simp also have "... = sem (cmd_double j) (execute (tm_double j) (0, tps) 0)" using tm_double_def exe_lt_length by simp also have "... = sem (cmd_double j) (0, tps)" by simp also have "... = (0, tps [j := tps ! j |:=| tosym (todigit 1) |+| 1, k := \xs ! (1 - 1)\])" using assms(7,8) sem_cmd_double_0[OF assms(1-2) _ _ assms(4,5,6)] by simp also have "... = (0, tps [j := (\\ # take (1 - 1) xs @ drop 1 xs\, Suc 1), k := \xs ! (1 - 1)\])" proof - have "tps ! j |:=| tosym (todigit 1) |+| 1 = (\xs\, 1) |:=| tosym (todigit 1) |+| 1" using assms(5) by simp also have "... = (\xs\(1 := tosym (todigit 1)), Suc 1)" by simp also have "... = (\xs\(1 := \), Suc 1)" by auto also have "... = (\\ # drop 1 xs\, Suc 1)" proof - have "\\ # drop 1 xs\ = \xs\(1 := \)" proof fix i :: nat consider "i = 0" | "i = 1" | "i > 1 \ i \ length xs" | "i > length xs" by linarith then show "\\ # drop 1 xs\ i = (\xs\(1 := \)) i" proof (cases) case 1 then show ?thesis by simp next case 2 then show ?thesis by simp next case 3 then have "\\ # drop 1 xs\ i = (\ # drop 1 xs) ! (i - 1)" using assms(3) by simp also have "... = (drop 1 xs) ! (i - 2)" using 3 by (metis Suc_1 diff_Suc_eq_diff_pred nth_Cons_pos zero_less_diff) also have "... = xs ! (Suc (i - 2))" using 3 assms(5) by simp also have "... = xs ! (i - 1)" using 3 by (metis Suc_1 Suc_diff_Suc) also have "... = \xs\ i" using 3 by simp also have "... = (\xs\(1 := \)) i" using 3 by simp finally show ?thesis . next case 4 then show ?thesis by simp qed qed then show ?thesis by simp qed also have "... = (\\ # take (1 - 1) xs @ drop 1 xs\, Suc 1)" by simp finally show ?thesis by auto qed finally show ?case . next case (Suc t) let ?xs = "\ # take (t - 1) xs @ drop t xs" let ?z = "xs ! (t - 1)" let ?tps = "tps [j := (\?xs\, Suc t), k := \?z\]" have lenxs: "length ?xs = length xs" using Suc by simp have 0: "?xs ! t = xs ! t" proof - have "t > 0" using Suc by simp then have "length (\ # take (t - 1) xs) = t" using Suc by simp moreover have "length (drop t xs) > 0" using Suc by simp moreover have "drop t xs ! 0 = xs ! t" using Suc by simp ultimately have "((\ # take (t - 1) xs) @ drop t xs) ! t = xs ! t" by (metis diff_self_eq_0 less_not_refl3 nth_append) then show ?thesis by simp qed have 1: "bit_symbols ?xs" proof - have "bit_symbols (take (t - 1) xs)" using assms(2) by simp moreover have "bit_symbols (drop t xs)" using assms(2) by simp moreover have "bit_symbols [\]" by simp ultimately have "bit_symbols ([\] @ take (t - 1) xs @ drop t xs)" using bit_symbols_append by presburger then show ?thesis by simp qed have 2: "Suc t \ length ?xs" using Suc by simp have 3: "Suc t > 0" using Suc by simp have 4: "length ?tps = Suc k" using assms by simp have 5: "?tps ! j = (\?xs\, Suc t)" by (simp add: Suc_lessD assms(1,4) nat_neq_iff) have 6: "?tps ! k = \?z\" by (simp add: assms(4)) have "execute (tm_double j) (0, tps) (Suc t) = exe (tm_double j) (execute (tm_double j) (0, tps) t)" by simp also have "... = sem (cmd_double j) (execute (tm_double j) (0, tps) t)" using tm_double_def exe_lt_length Suc by simp also have "... = sem (cmd_double j) (0, ?tps)" using Suc by simp also have "... = (0, ?tps [j := ?tps ! j |:=| tosym (todigit ?z) |+| 1, k := \?xs ! (Suc t - 1)\])" using sem_cmd_double_0[OF assms(1) 1 2 3 4 5 6] by simp also have "... = (0, ?tps [j := ?tps ! j |:=| tosym (todigit ?z) |+| 1, k := \xs ! (Suc t - 1)\])" using 0 by simp also have "... = (0, tps [j := ?tps ! j |:=| tosym (todigit ?z) |+| 1, k := \xs ! (Suc t - 1)\])" - using assms by (smt (z3) list_update_overwrite list_update_swap) + using assms by (smt (verit) list_update_overwrite list_update_swap) also have "... = (0, tps [j := (\?xs\, Suc t) |:=| tosym (todigit ?z) |+| 1, k := \xs ! (Suc t - 1)\])" using 5 by simp also have "... = (0, tps [j := (\?xs\(Suc t := tosym (todigit ?z)), Suc (Suc t)), k := \xs ! (Suc t - 1)\])" by simp also have "... = (0, tps [j := (\2 # take (Suc t - 1) xs @ drop (Suc t) xs\, Suc (Suc t)), k := \xs ! (Suc t - 1)\])" proof - have "\?xs\(Suc t := tosym (todigit ?z)) = \\ # take (Suc t - 1) xs @ drop (Suc t) xs\" proof fix i :: nat consider "i = 0" | "i > 0 \ i < Suc t" | "i = Suc t" | "i > Suc t \ i \ length xs" | "i > length xs" by linarith then show "(\?xs\(Suc t := tosym (todigit ?z))) i = \\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i" proof (cases) case 1 then show ?thesis by simp next case 2 then have lhs: "(\?xs\(Suc t := tosym (todigit ?z))) i = ?xs ! (i - 1)" using lenxs Suc by simp have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = (\ # take (Suc t - 1) xs @ drop (Suc t) xs) ! (i - 1)" using Suc 2 by auto then have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = ((\ # take (Suc t - 1) xs) @ drop (Suc t) xs) ! (i - 1)" by simp moreover have "length (\ # take (Suc t - 1) xs) = Suc t" using Suc.prems by simp ultimately have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = (\ # take (Suc t - 1) xs) ! (i - 1)" using 2 by (metis Suc_diff_1 Suc_lessD nth_append) also have "... = (\ # take t xs) ! (i - 1)" by simp also have "... = (\ # take (t - 1) xs @ [xs ! (t - 1)]) ! (i - 1)" using Suc by (metis Suc_diff_le Suc_le_lessD Suc_lessD diff_Suc_1 take_Suc_conv_app_nth) also have "... = ((\ # take (t - 1) xs) @ [xs ! (t - 1)]) ! (i - 1)" by simp also have "... = (\ # take (t - 1) xs) ! (i - 1)" using 2 Suc by (metis One_nat_def Suc_leD Suc_le_eq Suc_pred length_Cons length_take less_Suc_eq_le min_absorb2 nth_append) also have "... = ((\ # take (t - 1) xs) @ drop t xs) ! (i - 1)" using 2 Suc by (metis Suc_diff_1 Suc_diff_le Suc_leD Suc_lessD diff_Suc_1 length_Cons length_take less_Suc_eq min_absorb2 nth_append) also have "... = ?xs ! (i - 1)" by simp finally have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = ?xs ! (i - 1)" . then show ?thesis using lhs by simp next case 3 moreover have "?z = \ \ ?z = \" using `bit_symbols ?xs` Suc assms(2) by (metis Suc_diff_le Suc_leD Suc_le_lessD diff_Suc_1) ultimately have lhs: "(\?xs\(Suc t := tosym (todigit ?z))) i = ?z" by auto have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = \(\ # take t xs) @ drop (Suc t) xs\ (Suc t)" using 3 by simp also have "... = ((\ # take t xs) @ drop (Suc t) xs) ! t" using 3 Suc by simp also have "... = (\ # take t xs) ! t" using Suc by (metis Suc_leD length_Cons length_take lessI min_absorb2 nth_append) also have "... = xs ! (t - 1)" using Suc by simp finally have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = ?z" . then show ?thesis using lhs by simp next case 4 then have "(\?xs\(Suc t := tosym (todigit ?z))) i = \?xs\ i" by simp also have "... = ?xs ! (i - 1)" using 4 by auto also have "... = ((\ # take (t - 1) xs) @ drop t xs) ! (i - 1)" by simp also have "... = drop t xs ! (i - 1 - t)" using 4 Suc by (smt (verit, ccfv_threshold) Cons_eq_appendI Suc_diff_1 Suc_leD add_diff_cancel_right' bot_nat_0.extremum_uniqueI diff_diff_cancel length_append length_drop lenxs not_le not_less_eq nth_append) also have "... = xs ! (i - 1)" using 4 Suc by simp finally have lhs: "(\?xs\(Suc t := tosym (todigit ?z))) i = xs ! (i - 1)" . have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = (\ # take (Suc t - 1) xs @ drop (Suc t) xs) ! (i - 1)" using 4 by auto also have "... = ((\ # take t xs) @ drop (Suc t) xs) ! (i - 1)" by simp also have "... = (drop (Suc t) xs) ! (i - 1 - Suc t)" using Suc 4 - by (smt (z3) Suc_diff_1 Suc_leD Suc_leI bot_nat_0.extremum_uniqueI length_Cons length_take + by (smt (verit) Suc_diff_1 Suc_leD Suc_leI bot_nat_0.extremum_uniqueI length_Cons length_take min_absorb2 not_le nth_append) also have "... = xs ! (i - 1)" using Suc 4 Suc_lessE by fastforce finally have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = xs ! (i - 1)" . then show ?thesis using lhs by simp next case 5 then have "(\?xs\(Suc t := tosym (todigit ?z))) i = \?xs\ i" using Suc by simp then have lhs: "(\?xs\(Suc t := tosym (todigit ?z))) i = \" using 5 contents_outofbounds lenxs by simp have "length (\ # take (Suc t - 1) xs @ drop (Suc t) xs) = length xs" using Suc by simp then have "\\ # take (Suc t - 1) xs @ drop (Suc t) xs\ i = \" using 5 contents_outofbounds by simp then show ?thesis using lhs by simp qed qed then show ?thesis by simp qed finally show ?case . qed lemma execute_tm_double_1: assumes "j < k" and "bit_symbols xs" and "length xs > 0" and "length tps = Suc k" and "tps ! j = (\xs\, 1)" and "tps ! k = \\\" shows "execute (tm_double j) (0, tps) (Suc (length xs)) = (1, tps [j := (\\ # xs\, length xs + 2), k := \\\])" proof - let ?z = "xs ! (length xs - 1)" let ?xs = "\ # take (length xs - 1) xs" have "?z \ \" using assms(2,3) by (metis One_nat_def Suc_1 diff_less less_Suc_eq not_less_eq numeral_3_eq_3) have z23: "?z = \ \ ?z = \" using assms(2,3) by (meson diff_less zero_less_one) have lenxs: "length ?xs = length xs" using assms(3) by (metis Suc_diff_1 diff_le_self length_Cons length_take min_absorb2) have 0: "bit_symbols ?xs" using assms(2) bit_symbols_append[of "[\]" "take (length xs - 1) xs"] by simp have "execute (tm_double j) (0, tps) (length xs) = (0, tps [j := (\\ # take (length xs - 1) xs @ drop (length xs) xs\, Suc (length xs)), k := \?z\])" using execute_tm_double_0[OF assms(1-6), where ?t="length xs"] assms(3) by simp then have *: "execute (tm_double j) (0, tps) (length xs) = (0, tps [j := (\?xs\, Suc (length ?xs)), k := \?z\])" (is "_ = (0, ?tps)") using lenxs by simp let ?i = "Suc (length ?xs)" have 1: "?i > length ?xs" by simp have 2: "length ?tps = Suc k" using assms(4) by simp have 3: "?tps ! j = (\?xs\, ?i)" using assms(1,4) by simp have 4: "?tps ! k = \?z\" using assms(4) by simp have "execute (tm_double j) (0, tps) (Suc (length xs)) = exe (tm_double j) (0, ?tps)" using * by simp also have "... = sem (cmd_double j) (0, ?tps)" using tm_double_def exe_lt_length by simp also have "... = (1, ?tps [j := ?tps ! j |:=| (if ?z = \ then \ else tosym (todigit ?z)) |+| 1, k := \\\])" using sem_cmd_double_1[OF assms(1) 0 1 2 3 4] by simp also have "... = (1, ?tps [j := ?tps ! j |:=| (tosym (todigit ?z)) |+| 1, k := \\\])" using `?z \ 1` by simp also have "... = (1, ?tps [j := (\?xs\, Suc (length ?xs)) |:=| (tosym (todigit ?z)) |+| 1, k := \\\])" using 3 by simp also have "... = (1, ?tps [j := (\?xs\, Suc (length ?xs)) |:=| ?z |+| 1, k := \\\])" using z23 One_nat_def Suc_1 add_2_eq_Suc' numeral_3_eq_3 by presburger also have "... = (1, tps [j := (\?xs\, Suc (length ?xs)) |:=| ?z |+| 1, k := \\\])" - by (smt (z3) list_update_overwrite list_update_swap) + by (smt (verit) list_update_overwrite list_update_swap) also have "... = (1, tps [j := (\?xs\(Suc (length ?xs) := ?z), length ?xs + 2), k := \\\])" by simp also have "... = (1, tps [j := (\?xs\(Suc (length ?xs) := ?z), length xs + 2), k := \\\])" using lenxs by simp also have "... = (1, tps [j := (\\ # xs\, length xs + 2), k := \\\])" proof - have "\?xs\(Suc (length ?xs) := ?z) = \\ # xs\" proof fix i consider "i = 0" | "i > 0 \ i \ length xs" | "i = Suc (length xs)" | "i > Suc (length xs)" by linarith then show "(\?xs\(Suc (length ?xs) := ?z)) i = \\ # xs\ i" proof (cases) case 1 then show ?thesis by simp next case 2 then have "(\?xs\(Suc (length ?xs) := ?z)) i = \?xs\ i" using lenxs by simp also have "... = ?xs ! (i - 1)" using 2 by auto also have "... = (\ # xs) ! (i - 1)" using lenxs 2 assms(3) by (metis Suc_diff_1 Suc_le_lessD nth_take take_Suc_Cons) also have "... = \\ # xs\ i" using 2 by simp finally show ?thesis . next case 3 then have lhs: "(\?xs\(Suc (length ?xs) := ?z)) i = ?z" using lenxs by simp have "\\ # xs\ i = (\ # xs) ! (i - 1)" using 3 lenxs by simp also have "... = xs ! (i - 2)" using 3 assms(3) by simp also have "... = ?z" using 3 by simp finally have "\\ # xs\ i = ?z" . then show ?thesis using lhs by simp next case 4 then show ?thesis using 4 lenxs by simp qed qed then show ?thesis by simp qed finally show ?thesis . qed lemma execute_tm_double_Nil: assumes "j < k" and "length tps = Suc k" and "tps ! j = (\[]\, 1)" and "tps ! k = \\\" shows "execute (tm_double j) (0, tps) (Suc 0) = (1, tps [j := (\[]\, 2), k := \\\])" proof - have "execute (tm_double j) (0, tps) (Suc 0) = exe (tm_double j) (execute (tm_double j) (0, tps) 0)" by simp also have "... = exe (tm_double j) (0, tps)" by simp also have "... = sem (cmd_double j) (0, tps)" using tm_double_def exe_lt_length by simp also have "... = (1, tps [j := tps ! j |:=| (if (1::nat) = 1 then 0 else tosym (todigit 1)) |+| 1, k := \\\])" using sem_cmd_double_1[OF assms(1) _ _ assms(2-4)] by simp also have "... = (1, tps [j := tps ! j |:=| \ |+| 1, k := \\\])" by simp also have "... = (1, tps [j := (\[]\, 1) |:=| \ |+| 1, k := \\\])" using assms(3) by simp also have "... = (1, tps [j := (\[]\(1 := \), 2), k := \\\])" by (metis fst_eqD one_add_one snd_eqD) also have "... = (1, tps [j := (\[]\, 2), k := \\\])" by (metis contents_outofbounds fun_upd_idem_iff list.size(3) zero_less_one) finally show ?thesis . qed lemma execute_tm_double: assumes "j < k" and "length tps = Suc k" and "tps ! j = (\canrepr n\, 1)" and "tps ! k = \\\" shows "execute (tm_double j) (0, tps) (Suc (length (canrepr n))) = (1, tps [j := (\canrepr (2 * n)\, length (canrepr n) + 2), k := \\\])" proof (cases "n = 0") case True then have "canrepr n = []" using canrepr_0 by simp then show ?thesis using execute_tm_double_Nil[OF assms(1-2) _ assms(4)] assms(3) True by (metis add_2_eq_Suc' list.size(3) mult_0_right numeral_2_eq_2) next case False let ?xs = "canrepr n" have "num (\ # ?xs) = 2 * num ?xs" using num_Cons by simp then have "num (\ # ?xs) = 2 * n" using canrepr by simp moreover have "canonical (\ # ?xs)" proof - have "?xs \ []" using False canrepr canrepr_0 by metis then show ?thesis using canonical_Cons canonical_canrepr by simp qed ultimately have "canrepr (2 * n) = \ # ?xs" using canreprI by blast then show ?thesis using execute_tm_double_1[OF assms(1) _ _ assms(2) _ assms(4)] assms(3) False canrepr canrepr_0 bit_symbols_canrepr by (metis length_greater_0_conv) qed lemma execute_tm_double_app: assumes "j < k" and "length tps = k" and "tps ! j = (\canrepr n\, 1)" shows "execute (tm_double j) (0, tps @ [\\\]) (Suc (length (canrepr n))) = (1, tps [j := (\canrepr (2 * n)\, length (canrepr n) + 2)] @ [\\\])" proof - let ?tps = "tps @ [\\\]" have "length ?tps = Suc k" using assms(2) by simp moreover have "?tps ! j = (\canrepr n\, 1)" using assms(1,2,3) by (simp add: nth_append) moreover have "?tps ! k = \\\" using assms(2) by (simp add: nth_append) moreover have "tps [j := (\canrepr (2 * n)\, length (canrepr n) + 2)] @ [\\\] = ?tps [j := (\canrepr (2 * n)\, length (canrepr n) + 2), k := \\\]" using assms by (metis length_list_update list_update_append1 list_update_length) ultimately show ?thesis using assms execute_tm_double[OF assms(1), where ?tps="tps @ [\\\]"] by simp qed lemma transforms_tm_double: assumes "j < k" and "length tps = k" and "tps ! j = (\canrepr n\, 1)" shows "transforms (tm_double j) (tps @ [\\\]) (Suc (length (canrepr n))) (tps [j := (\canrepr (2 * n)\, length (canrepr n) + 2)] @ [\\\])" using assms transforms_def transits_def tm_double_def execute_tm_double_app by auto lemma tm_double_immobile: fixes k :: nat assumes "j > 0" and "j < k" shows "immobile (tm_double j) k (Suc k)" proof - let ?M = "tm_double j" { fix q :: nat and rs :: "symbol list" assume q: "q < length ?M" assume rs: "length rs = Suc k" then have len: "length rs - 1 = k" by simp have neq: "k \ j" using assms(2) by simp have "?M ! q = cmd_double j" using tm_double_def q by simp moreover have "(cmd_double j) rs [!] k = (tosym (todigit (rs ! j)), Stay)" using cmd_double_def rs len neq by fastforce ultimately have "(cmd_double j) rs [~] k = Stay" by simp } then show ?thesis by (simp add: immobile_def tm_double_def) qed lemma tm_double_bounded_write: assumes "j < k - 1" shows "bounded_write (tm_double j) (k - 1) 4" using assms cmd_double_def tm_double_def bounded_write_def by simp text \ The next Turing machine removes the memorization tape. \ definition tm_double' :: "nat \ machine" where "tm_double' j \ cartesian (tm_double j) 4" lemma tm_double'_tm: assumes "j > 0" and "k \ 2" and "G \ 4" and "j < k" shows "turing_machine k G (tm_double' j)" unfolding tm_double'_def using assms cartesian_tm tm_double_tm by simp lemma transforms_tm_double'I [transforms_intros]: assumes "j > 0" and "j < k" and "length tps = k" and "tps ! j = (\canrepr n\, 1)" and "t = (Suc (length (canrepr n)))" and "tps' = tps [j := (\canrepr (2 * n)\, length (canrepr n) + 2)]" shows "transforms (tm_double' j) tps t tps'" unfolding tm_double'_def proof (rule cartesian_transforms_onesie) show "turing_machine (Suc k) 4 (tm_double j)" using assms(1,2) tm_double_tm by simp show "length tps = k" "2 \ k" "(1::nat) < 4" using assms by simp_all show "bounded_write (tm_double j) k 4" by (metis assms(2) diff_Suc_1 tm_double_bounded_write) show "immobile (tm_double j) k (Suc k)" by (simp add: assms(1,2) tm_double_immobile) show "transforms (tm_double j) (tps @ [\\\]) t (tps' @ [\\\])" using assms transforms_tm_double by simp qed text \ The next Turing machine is the one we actually use to double a number. It runs @{const tm_double'} and performs a carriage return. \ definition tm_times2 :: "tapeidx \ machine" where "tm_times2 j \ tm_double' j ;; tm_cr j" lemma tm_times2_tm: assumes "k \ 2" and "j > 0" and "j < k" and "G \ 4" shows "turing_machine k G (tm_times2 j)" using assms by (simp add: assms(1) tm_cr_tm tm_double'_tm tm_times2_def) lemma transforms_tm_times2I [transforms_intros]: assumes "j > 0" and "j < k" and "length tps = k" and "tps ! j = (\n\\<^sub>N, 1)" and "t = 5 + 2 * nlength n" and "tps' = tps [j := (\2 * n\\<^sub>N, 1)]" shows "transforms (tm_times2 j) tps t tps'" unfolding tm_times2_def proof (tform tps: assms) show "clean_tape (tps[j := (\2 * n\\<^sub>N, nlength n + 2)] ! j)" using clean_tape_ncontents assms by simp show "t = Suc (nlength n) + (tps[j := (\2 * n\\<^sub>N, nlength n + 2)] :#: j + 2)" using assms by simp qed subsubsection \Multiplying arbitrary numbers\ text \ Before we can multiply arbitrary numbers we need just a few more lemmas. \null \ lemma num_drop_le_nu: "num (drop j xs) \ num xs" proof (cases "j \ length xs") case True let ?ys = "drop j xs" have map_shift_upt: "map (\i. f (j + i)) [0.. nat" and j l by (rule nth_equalityI) simp_all have "num ?ys = (\i\[0..i\[0.. 2 ^ j * (\i\[0..i\[0..i\[0..i\[j..i. todigit (xs ! i) * 2 ^ i" j "length ?ys"] by simp also have "... \ (\i\[0..i\[j..i\[0..i\[0.. length xs" shows "nlength (prod xs ys i) \ nlength (num xs * num ys)" using prod[OF assms] num_drop_le_nu mult_le_mono1 nlength_mono by simp corollary nlength_prod'_le_prod: assumes "i \ nlength x" shows "nlength (prod' x y i) \ nlength (x * y)" using assms prod'_def nlength_prod_le_prod by (metis prod' prod_eq_prod) lemma two_times_prod: assumes "i < length xs" shows "2 * prod xs ys i \ num xs * num ys" proof - have "2 * prod xs ys i \ prod xs ys (Suc i)" by simp also have "... = num (drop (length xs - Suc i) xs) * num ys" using prod[of "Suc i" xs] assms by simp also have "... \ num xs * num ys" using num_drop_le_nu by simp finally show ?thesis . qed corollary two_times_prod': assumes "i < nlength x" shows "2 * prod' x y i \ x * y" using assms two_times_prod prod'_def by (metis prod' prod_eq_prod) text \ The next Turing machine multiplies the numbers on tapes $j_1$ and $j_2$ and writes the result to tape $j_3$. It iterates over the binary digits on $j_1$ starting from the most significant digit. In each iteration it doubles the intermediate result on $j_3$. If the current digit is @{text \}, the number on $j_2$ is added to $j_3$. \ definition tm_mult :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_mult j1 j2 j3 \ tm_right_until j1 {\} ;; tm_left j1 ;; WHILE [] ; \rs. rs ! j1 \ \ DO tm_times2 j3 ;; IF \rs. rs ! j1 = \ THEN tm_add j2 j3 ELSE [] ENDIF ;; tm_left j1 DONE ;; tm_right j1" lemma tm_mult_tm: assumes "j1 \ j2" "j2 \ j3" "j3 \ j1" and "j3 > 0" assumes "k \ 2" and "G \ 4" and "j1 < k" "j2 < k" "j3 < k" shows "turing_machine k G (tm_mult j1 j2 j3)" unfolding tm_mult_def using assms tm_left_tm tm_right_tm Nil_tm tm_add_tm tm_times2_tm tm_right_until_tm turing_machine_branch_turing_machine turing_machine_loop_turing_machine by simp locale turing_machine_mult = fixes j1 j2 j3 :: tapeidx begin definition "tm1 \ tm_right_until j1 {\}" definition "tm2 \ tm1 ;; tm_left j1" definition "tmIf \ IF \rs. rs ! j1 = \ THEN tm_add j2 j3 ELSE [] ENDIF" definition "tmBody1 \ tm_times2 j3 ;; tmIf" definition "tmBody \ tmBody1 ;; tm_left j1" definition "tmWhile \ WHILE [] ; \rs. rs ! j1 \ \ DO tmBody DONE" definition "tm3 \ tm2 ;; tmWhile" definition "tm4 \ tm3 ;; tm_right j1" lemma tm4_eq_tm_mult: "tm4 = tm_mult j1 j2 j3" using tm1_def tm2_def tm3_def tm4_def tm_mult_def tmIf_def tmBody_def tmBody1_def tmWhile_def by simp context fixes x y k :: nat and tps0 :: "tape list" assumes jk: "j1 \ j2" "j2 \ j3" "j3 \ j1" "j3 > 0" "j1 < k" "j2 < k" "j3 < k" "length tps0 = k" assumes tps0: "tps0 ! j1 = (\x\\<^sub>N, 1)" "tps0 ! j2 = (\y\\<^sub>N, 1)" "tps0 ! j3 = (\0\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j1 := (\x\\<^sub>N, Suc (nlength x))]" lemma tm1 [transforms_intros]: assumes "t = Suc (nlength x)" shows "transforms tm1 tps0 t tps1" unfolding tm1_def proof (tform tps: assms tps0 tps1_def jk) show "rneigh (tps0 ! j1) {\} (nlength x)" proof (rule rneighI) show "(tps0 ::: j1) (tps0 :#: j1 + nlength x) \ {\}" by (simp add: tps0) show "\n'. n' < nlength x \ (tps0 ::: j1) (tps0 :#: j1 + n') \ {\}" using tps0 bit_symbols_canrepr contents_def by fastforce qed qed definition "tps2 \ tps0 [j1 := (\x\\<^sub>N, nlength x)]" lemma tm2 [transforms_intros]: assumes "t = Suc (Suc (nlength x))" and "tps' = tps2" shows "transforms tm2 tps0 t tps'" unfolding tm2_def by (tform tps: assms tps1_def tps2_def jk) definition "tpsL t \ tps0 [j1 := (\x\\<^sub>N, nlength x - t), j3 := (\prod' x y t\\<^sub>N, 1)]" definition "tpsL1 t \ tps0 [j1 := (\x\\<^sub>N, nlength x - t), j3 := (\2 * prod' x y t\\<^sub>N, 1)]" definition "tpsL2 t \ tps0 [j1 := (\x\\<^sub>N, nlength x - t), j3 := (\prod' x y (Suc t)\\<^sub>N, 1)]" definition "tpsL3 t \ tps0 [j1 := (\x\\<^sub>N, nlength x - t - 1), j3 := (\prod' x y (Suc t)\\<^sub>N, 1)]" lemma tmIf [transforms_intros]: assumes "t < nlength x" and "ttt = 12 + 3 * nlength (x * y)" shows "transforms tmIf (tpsL1 t) ttt (tpsL2 t)" unfolding tmIf_def proof (tform tps: assms tpsL1_def tps0 jk) have "nlength y \ nlength (x * y) \ nlength (2 * prod' x y t) \ nlength (x * y)" proof have "x > 0" using assms(1) gr_implies_not_zero nlength_0 by auto then have "y \ x * y" by simp then show "nlength y \ nlength (x * y)" using nlength_mono by simp show "nlength (2 * prod' x y t) \ nlength (x * y)" using assms(1) by (simp add: nlength_mono two_times_prod') qed then show "3 * max (nlength y) (nlength (2 * Arithmetic.prod' x y t)) + 10 + 2 \ ttt" using assms(2) by simp let ?xs = "canrepr x" and ?ys = "canrepr y" let ?r = "read (tpsL1 t) ! j1" have "?r = (\x\\<^sub>N) (nlength x - t)" using tpsL1_def jk tapes_at_read' by (metis fst_conv length_list_update list_update_swap nth_list_update_eq snd_conv) then have r: "?r = canrepr x ! (nlength x - 1 - t)" using assms contents_def by simp have "prod' x y (Suc t) = 2 * prod' x y t + (if ?xs ! (length ?xs - 1 - t) = \ then num ?ys else 0)" using prod'_def by simp also have "... = 2 * prod' x y t + (if ?r = \ then num ?ys else 0)" using r by simp also have "... = 2 * prod' x y t + (if ?r = \ then y else 0)" using canrepr by simp finally have "prod' x y (Suc t) = 2 * prod' x y t + (if ?r = \ then y else 0)" . then show "read (tpsL1 t) ! j1 \ \ \ tpsL2 t = tpsL1 t" and "read (tpsL1 t) ! j1 = \ \ tpsL2 t = (tpsL1 t) [j3 := (\y + 2 * Arithmetic.prod' x y t\\<^sub>N, 1)]" by (simp_all add: add.commute tpsL1_def tpsL2_def) qed lemma tmBody1 [transforms_intros]: assumes "t < nlength x" and "ttt = 17 + 2 * nlength (Arithmetic.prod' x y t) + 3 * nlength (x * y)" shows "transforms tmBody1 (tpsL t) ttt (tpsL2 t)" unfolding tmBody1_def by (tform tps: jk tpsL_def tpsL1_def assms(1) time: assms(2)) lemma tmBody: assumes "t < nlength x" and "ttt = 6 + 2 * nlength (prod' x y t) + (12 + 3 * nlength (x * y))" shows "transforms tmBody (tpsL t) ttt (tpsL (Suc t))" unfolding tmBody_def by (tform tps: jk tpsL_def tpsL2_def assms(1) time: assms(2)) lemma tmBody' [transforms_intros]: assumes "t < nlength x" and "ttt = 18 + 5 * nlength (x * y)" shows "transforms tmBody (tpsL t) ttt (tpsL (Suc t))" proof - have "6 + 2 * nlength (prod' x y t) + (12 + 3 * nlength (x * y)) \ 18 + 5 * nlength (x * y)" using assms nlength_prod'_le_prod by simp then show ?thesis using tmBody assms transforms_monotone by blast qed lemma read_contents: fixes tps :: "tape list" and j :: tapeidx and zs :: "symbol list" assumes "tps ! j = (\zs\, i)" and "i > 0" and "i \ length zs" and "j < length tps" shows "read tps ! j = zs ! (i - 1)" using assms tapes_at_read' by fastforce lemma tmWhile [transforms_intros]: assumes "ttt = 1 + 25 * (nlength x + nlength y) * (nlength x + nlength y)" shows "transforms tmWhile (tpsL 0) ttt (tpsL (nlength x))" unfolding tmWhile_def proof (tform) show "read (tpsL i) ! j1 \ \" if "i < nlength x" for i proof - have "(tpsL i) ! j1 = (\x\\<^sub>N, nlength x - i)" using tpsL_def jk by simp moreover have *: "nlength x - i > 0" "nlength x - i \ length (canrepr x)" using that by simp_all moreover have "length (tpsL i) = k" using tpsL_def jk by simp ultimately have "read (tpsL i) ! j1 = canrepr x ! (nlength x - i - 1)" using jk read_contents by simp then show ?thesis using * bit_symbols_canrepr by (metis One_nat_def Suc_le_lessD Suc_pred less_numeral_extra(4) proper_symbols_canrepr) qed show "\ read (tpsL (nlength x)) ! j1 \ \" proof - have "(tpsL (nlength x)) ! j1 = (\x\\<^sub>N, nlength x - nlength x)" using tpsL_def jk by simp then have "(tpsL (nlength x)) ! j1 = (\x\\<^sub>N, 0)" by simp then have "read (tpsL (nlength x)) ! j1 = \" using tapes_at_read' tpsL_def contents_at_0 jk by (metis fst_conv length_list_update snd_conv) then show ?thesis by simp qed show "nlength x * (18 + 5 * nlength (x * y) + 2) + 1 \ ttt" proof (cases "x = 0") case True then show ?thesis using assms by simp next case False have "nlength x * (18 + 5 * nlength (x * y) + 2) + 1 = nlength x * (20 + 5 * nlength (x * y)) + 1" by simp also have "... \ nlength x * (20 + 5 * (nlength x + nlength y)) + 1" using nlength_prod by (meson add_mono le_refl mult_le_mono) also have "... \ nlength x * (20 * (nlength x + nlength y) + 5 * (nlength x + nlength y)) + 1" proof - have "1 \ nlength x + nlength y" using False nlength_0 by (simp add: Suc_leI) then show ?thesis by simp qed also have "... \ nlength x * (25 * (nlength x + nlength y)) + 1" by simp also have "... \ (nlength x + nlength y) * (25 * (nlength x + nlength y)) + 1" by simp finally show ?thesis using assms by linarith qed qed lemma tm3: assumes "ttt = Suc (Suc (nlength x)) + Suc ((25 * nlength x + 25 * nlength y) * (nlength x + nlength y))" shows "transforms tm3 tps0 ttt (tpsL (nlength x))" unfolding tm3_def proof (tform time: assms) show "tpsL 0 = tps2" proof - have "prod' x y 0 = 0" using prod'_def by simp then show ?thesis using tpsL_def tps2_def jk tps0 by (metis diff_zero list_update_id list_update_swap) qed qed definition "tps3 \ tps0 [j1 := (\x\\<^sub>N, 0), j3 := (\x * y\\<^sub>N, 1)]" lemma tm3' [transforms_intros]: assumes "ttt = 3 + 26 * (nlength x + nlength y) * (nlength x + nlength y)" shows "transforms tm3 tps0 ttt tps3" proof - have "Suc (Suc (nlength x)) + Suc ((25 * nlength x + 25 * nlength y) * (nlength x + nlength y)) \ Suc (Suc (nlength x + nlength y)) + Suc ((25 * nlength x + 25 * nlength y) * (nlength x + nlength y))" by simp also have "... \ 2 + (nlength x + nlength y) * (nlength x + nlength y) + 1 + 25 * (nlength x + nlength y) * (nlength x + nlength y)" by (simp add: le_square) also have "... = 3 + 26 * (nlength x + nlength y) * (nlength x + nlength y)" by linarith finally have "Suc (Suc (nlength x)) + Suc ((25 * nlength x + 25 * nlength y) * (nlength x + nlength y)) \ 3 + 26 * (nlength x + nlength y) * (nlength x + nlength y)" . moreover have "tps3 = tpsL (nlength x)" using tps3_def tpsL_def by (simp add: prod') ultimately show ?thesis using tm3 assms transforms_monotone by simp qed definition "tps4 \ tps0 [j3 := (\x * y\\<^sub>N, 1)]" lemma tm4: assumes "ttt = 4 + 26 * (nlength x + nlength y) * (nlength x + nlength y)" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: tps3_def jk time: assms) show "tps4 = tps3[j1 := tps3 ! j1 |+| 1]" using tps4_def tps3_def jk tps0 by (metis One_nat_def add.right_neutral add_Suc_right fst_conv list_update_id list_update_overwrite list_update_swap nth_list_update_eq nth_list_update_neq snd_conv) qed end (* context x y k tps0 *) end (* locale turing_machine_mult *) lemma transforms_tm_mult [transforms_intros]: fixes j1 j2 j3 :: tapeidx and x y k ttt :: nat and tps tps' :: "tape list" assumes "j1 \ j2" "j2 \ j3" "j3 \ j1" "j3 > 0" assumes "length tps = k" and "j1 < k" "j2 < k" "j3 < k" and "tps ! j1 = (\x\\<^sub>N, 1)" and "tps ! j2 = (\y\\<^sub>N, 1)" and "tps ! j3 = (\0\\<^sub>N, 1)" and "ttt = 4 + 26 * (nlength x + nlength y) * (nlength x + nlength y)" and "tps' = tps [j3 := (\x * y\\<^sub>N, 1)]" shows "transforms (tm_mult j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_mult j1 j2 j3 . show ?thesis using assms loc.tps4_def loc.tm4 loc.tm4_eq_tm_mult by metis qed subsection \Powers\ text \ In this section we construct for every $d \in \nat$ a Turing machine that computes $n^d$. The following TMs expect a number $n$ on tape $j_1$ and output $n^d$ on tape $j_3$. Another tape, $j_2$, is used as scratch space to hold intermediate values. The TMs initialize tape $j_3$ with~1 and then multiply this value by $n$ for $d$ times using the TM @{const tm_mult}. \ fun tm_pow :: "nat \ tapeidx \ tapeidx \ tapeidx \ machine" where "tm_pow 0 j1 j2 j3 = tm_setn j3 1" | "tm_pow (Suc d) j1 j2 j3 = tm_pow d j1 j2 j3 ;; (tm_copyn j3 j2 ;; tm_setn j3 0 ;; tm_mult j1 j2 j3 ;; tm_setn j2 0)" lemma tm_pow_tm: assumes "j1 \ j2" "j2 \ j3" "j3 \ j1" and "0 < j2" "0 < j3" "0 < j1" assumes "j1 < k" "j2 < k" "j3 < k" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_pow d j1 j2 j3)" using assms tm_copyn_tm tm_setn_tm tm_mult_tm by (induction d) simp_all locale turing_machine_pow = fixes j1 j2 j3 :: tapeidx begin definition "tm1 \ tm_copyn j3 j2 ;; tm_setn j3 0" definition "tm2 \ tm1 ;; tm_mult j1 j2 j3" definition "tm3 \ tm2 ;; tm_setn j2 0" fun tm4 :: "nat \ machine" where "tm4 0 = tm_setn j3 1" | "tm4 (Suc d) = tm4 d ;; tm3" lemma tm4_eq_tm_pow: "tm4 d = tm_pow d j1 j2 j3" using tm3_def tm2_def tm1_def by (induction d) simp_all context fixes x y k :: nat and tps0 :: "tape list" assumes jk: "k = length tps0" "j1 < k" "j2 < k" "j3 < k" "j1 \ j2" "j2 \ j3" "j3 \ j1" "0 < j2" "0 < j3" "0 < j1" assumes tps0: "tps0 ! j1 = (\x\\<^sub>N, 1)" "tps0 ! j2 = (\0\\<^sub>N, 1)" "tps0 ! j3 = (\y\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j2 := (\y\\<^sub>N, 1), j3 := (\0\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 24 + 5 * nlength y" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: assms jk tps0 tps1_def) show "ttt = 14 + 3 * (nlength y + nlength 0) + (10 + 2 * nlength y + 2 * nlength 0)" using assms by simp qed definition "tps2 \ tps0 [j2 := (\y\\<^sub>N, 1), j3 := (\x * y\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 28 + 5 * nlength y + (26 * nlength x + 26 * nlength y) * (nlength x + nlength y)" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: jk tps1_def time: assms) show "tps1 ! j1 = (\x\\<^sub>N, 1)" using jk tps0 tps1_def by simp show "tps2 = tps1[j3 := (\x * y\\<^sub>N, 1)]" using tps2_def tps1_def by simp qed definition "tps3 \ tps0 [j3 := (\x * y\\<^sub>N, 1)]" lemma tm3: assumes "ttt = 38 + 7 * nlength y + (26 * nlength x + 26 * nlength y) * (nlength x + nlength y)" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: jk tps2_def time: assms) show "tps3 = tps2[j2 := (\0\\<^sub>N, 1)]" using tps3_def tps2_def jk by (metis list_update_id list_update_overwrite list_update_swap tps0(2)) qed lemma tm3': assumes "ttt = 38 + 33 * (nlength x + nlength y) ^ 2" shows "transforms tm3 tps0 ttt tps3" proof - have "38 + 7 * nlength y + (26 * nlength x + 26 * nlength y) * (nlength x + nlength y) = 38 + 7 * nlength y + 26 * (nlength x + nlength y) * (nlength x + nlength y)" by simp also have "... \ 38 + 33 * (nlength x + nlength y) * (nlength x + nlength y)" proof - have "nlength y \ (nlength x + nlength y) * (nlength x + nlength y)" by (meson le_add2 le_square le_trans) then show ?thesis by linarith qed also have "... = 38 + 33 * (nlength x + nlength y) ^ 2" by algebra finally have "38 + 7 * nlength y + (26 * nlength x + 26 * nlength y) * (nlength x + nlength y) \ ttt" using assms(1) by simp then show ?thesis using tm3 transforms_monotone assms by meson qed end (* context x y k tps0 *) lemma tm3'' [transforms_intros]: fixes x d k :: nat and tps0 :: "tape list" assumes "k = length tps0" and "j1 < k" "j2 < k" "j3 < k" assumes j_neq [simp]: "j1 \ j2" "j2 \ j3" "j3 \ j1" and j_gt [simp]: "0 < j2" "0 < j3" "0 < j1" and "tps0 ! j1 = (\x\\<^sub>N, 1)" and "tps0 ! j2 = (\0\\<^sub>N, 1)" and "tps0 ! j3 = (\x ^ d\\<^sub>N, 1)" and "ttt = 71 + 99 * (Suc d) ^ 2 * (nlength x) ^ 2" and "tps' = tps0 [j3 := (\x ^ Suc d\\<^sub>N, 1)]" shows "transforms tm3 tps0 ttt tps'" proof - let ?l = "nlength x" have "transforms tm3 tps0 (38 + 33 * (nlength x + nlength (x ^ d)) ^ 2) tps'" using tm3' assms tps3_def by simp moreover have "38 + 33 * (nlength x + nlength (x ^ d)) ^ 2 \ 71 + 99 * (Suc d) ^ 2 * ?l ^ 2" proof - have "38 + 33 * (nlength x + nlength (x ^ d)) ^ 2 \ 38 + 33 * (Suc (Suc d * ?l)) ^ 2" using nlength_pow by simp also have "... = 38 + 33 * ((Suc d * ?l)^2 + 2 * (Suc d * ?l) * 1 + 1^2)" by (metis Suc_eq_plus1 add_Suc one_power2 power2_sum) also have "... = 38 + 33 * ((Suc d * ?l)^2 + 2 * (Suc d * ?l) + 1)" by simp also have "... \ 38 + 33 * ((Suc d * ?l)^2 + 2 * (Suc d * ?l)^2 + 1)" proof - have "(Suc d * ?l) \ (Suc d * ?l) ^ 2" by (simp add: le_square power2_eq_square) then show ?thesis by simp qed also have "... \ 38 + 33 * (3 * (Suc d * ?l)^2 + 1)" by simp also have "... = 38 + 33 * (3 * (Suc d) ^ 2 * ?l^2 + 1)" by algebra also have "... = 71 + 99 * (Suc d) ^ 2 * ?l ^ 2" by simp finally show ?thesis . qed ultimately show ?thesis using transforms_monotone assms(14) by blast qed context fixes x k :: nat and tps0 :: "tape list" assumes jk: "j1 < k" "j2 < k" "j3 < k" "j1 \ j2" "j2 \ j3" "j3 \ j1" "0 < j2" "0 < j3" "0 < j1" "k = length tps0" assumes tps0: "tps0 ! j1 = (\x\\<^sub>N, 1)" "tps0 ! j2 = (\0\\<^sub>N, 1)" "tps0 ! j3 = (\0\\<^sub>N, 1)" begin lemma tm4: fixes d :: nat assumes "tps' = tps0 [j3 := (\x ^ d\\<^sub>N, 1)]" and "ttt = 12 + 71 * d + 99 * d ^ 3 * (nlength x) ^ 2" shows "transforms (tm4 d) tps0 ttt tps'" using assms proof (induction d arbitrary: tps' ttt) case 0 have "tm4 0 = tm_setn j3 1" by simp let ?tps = "tps0 [j3 := (\1\\<^sub>N, 1)]" let ?t = "10 + 2 * nlength 1" have "transforms (tm_setn j3 1) tps0 ?t ?tps" using transforms_tm_setnI[of j3 tps0 0 ?t 1 ?tps] jk tps0 by simp then have "transforms (tm_setn j3 1) tps0 ?t tps'" using 0 by simp then show ?case using 0 nlength_1_simp by simp next case (Suc d) note Suc.IH [transforms_intros] let ?l = "nlength x" have "tm4 (Suc d) = tm4 d ;; tm3" by simp define t where "t = 12 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + (71 + 99 * (Suc d)\<^sup>2 * (nlength x)\<^sup>2)" have "transforms (tm4 d ;; tm3) tps0 t tps'" by (tform tps: jk tps0 Suc.prems(1) time: t_def) moreover have "t \ 12 + 71 * Suc d + 99 * Suc d ^ 3 * ?l\<^sup>2" proof - have "t = 12 + d * 71 + 99 * d ^ 3 * ?l\<^sup>2 + (71 + 99 * (Suc d)\<^sup>2 * ?l\<^sup>2)" using t_def by simp also have "... = 12 + Suc d * 71 + 99 * d ^ 3 * ?l\<^sup>2 + 99 * (Suc d)\<^sup>2 * ?l\<^sup>2" by simp also have "... = 12 + Suc d * 71 + 99 * ?l^2 * (d ^ 3 + (Suc d)\<^sup>2)" by algebra also have "... \ 12 + Suc d * 71 + 99 * ?l^2 * Suc d ^ 3" proof - have "Suc d ^ 3 = Suc d * Suc d ^ 2" by algebra also have "... = Suc d * (d ^ 2 + 2 * d + 1)" by (metis (no_types, lifting) Suc_1 add.commute add_Suc mult_2 one_power2 plus_1_eq_Suc power2_sum) also have "... = (d + 1) * (d ^ 2 + 2 * d + 1)" by simp also have "... = d ^ 3 + 2 * d ^ 2 + d + d ^ 2 + 2 * d + 1" by algebra also have "... = d ^ 3 + (d + 1) ^ 2 + 2 * d ^ 2 + d" by algebra also have "... \ d ^ 3 + (d + 1) ^ 2" by simp finally have "Suc d ^ 3 \ d ^ 3 + Suc d ^ 2" by simp then show ?thesis by simp qed also have "... = 12 + 71 * Suc d + 99 * Suc d ^ 3 * ?l^2" by simp finally show ?thesis . qed ultimately show ?case using transforms_monotone Suc by simp qed end (* context x k tps0 *) end (* locale turing_machine_power *) lemma transforms_tm_pow [transforms_intros]: fixes d :: nat assumes "j1 \ j2" "j2 \ j3" "j3 \ j1" "0 < j2" "0 < j3" "0 < j1" "j1 < k" "j2 < k" "j3 < k" "k = length tps" assumes "tps ! j1 = (\x\\<^sub>N, 1)" "tps ! j2 = (\0\\<^sub>N, 1)" "tps ! j3 = (\0\\<^sub>N, 1)" assumes "ttt = 12 + 71 * d + 99 * d ^ 3 * (nlength x) ^ 2" assumes "tps' = tps [j3 := (\x ^ d\\<^sub>N, 1)]" shows "transforms (tm_pow d j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_pow j1 j2 j3 . show ?thesis using assms loc.tm4_eq_tm_pow loc.tm4 by metis qed subsection \Monomials\ text \ A monomial is a power multiplied by a constant coefficient. The following Turing machines have parameters $c$ and $d$ and expect a number $x$ on tape $j$. They output $c\cdot x^d$ on tape $j + 3$. The tapes $j+1$ and $j+2$ are scratch space for use by @{const tm_pow} and @{const tm_mult}. \ definition tm_monomial :: "nat \ nat \ tapeidx \ machine" where "tm_monomial c d j \ tm_pow d j (j + 1) (j + 2) ;; tm_setn (j + 1) c ;; tm_mult (j + 1) (j + 2) (j + 3);; tm_setn (j + 1) 0 ;; tm_setn (j + 2) 0" lemma tm_monomial_tm: assumes "k \ 2" and "G \ 4" and "j + 3 < k" and "0 < j" shows "turing_machine k G (tm_monomial c d j)" unfolding tm_monomial_def using assms tm_setn_tm tm_mult_tm tm_pow_tm turing_machine_sequential_turing_machine by simp locale turing_machine_monomial = fixes c d :: nat and j :: tapeidx begin definition "tm1 \ tm_pow d j (j + 1) (j + 2)" definition "tm2 \ tm1 ;; tm_setn (j + 1) c" definition "tm3 \ tm2 ;; tm_mult (j + 1) (j + 2) (j + 3)" definition "tm4 \ tm3 ;; tm_setn (j + 1) 0" definition "tm5 \ tm4 ;; tm_setn (j + 2) 0" lemma tm5_eq_tm_monomial: "tm5 = tm_monomial c d j" unfolding tm1_def tm2_def tm3_def tm4_def tm5_def tm_monomial_def by simp context fixes x k :: nat and tps0 :: "tape list" assumes jk: "k = length tps0" "j + 3 < k" "0 < j" assumes tps0: "tps0 ! j = (\x\\<^sub>N, 1)" "tps0 ! (j + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j + 3) = (\0\\<^sub>N, 1)" begin definition "tps1 \ tps0 [(j + 2) := (\x ^ d\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 12 + 71 * d + 99 * d ^ 3 * (nlength x) ^ 2" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def by (tform tps: assms tps0 jk tps1_def) definition "tps2 \ tps0 [j + 2 := (\x ^ d\\<^sub>N, 1), j + 1 := (\c\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 22 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 2 * nlength c" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: assms tps0 jk tps2_def tps1_def) show "ttt = 12 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + (10 + 2 * nlength 0 + 2 * nlength c)" using assms(1) by simp qed definition "tps3 \ tps0 [j + 2 := (\x ^ d\\<^sub>N, 1), j + 1 := (\c\\<^sub>N, 1), j + 3 := (\c * x ^ d\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 26 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 2 * nlength c + 26 * (nlength c + nlength (x ^ d)) ^ 2" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps2_def tps3_def tps0 jk) show "ttt = 22 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 2 * nlength c + (4 + 26 * (nlength c + nlength (x ^ d)) * (nlength c + nlength (x ^ d)))" using assms by algebra qed definition "tps4 \ tps0 [j + 2 := (\x ^ d\\<^sub>N, 1), j + 3 := (\c * x ^ d\\<^sub>N, 1)]" lemma tm4 [transforms_intros]: assumes "ttt = 36 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 4 * nlength c + 26 * (nlength c + nlength (x ^ d))\<^sup>2" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: tps4_def tps3_def tps0 jk time: assms) show "tps4 = tps3[j + 1 := (\0\\<^sub>N, 1)]" unfolding tps4_def tps3_def using jk tps0(2) list_update_id[of tps0 "Suc j"] by (simp add: list_update_swap) qed definition "tps5 \ tps0 [j + 3 := (\c * x ^ d\\<^sub>N, 1)]" lemma tm5: assumes "ttt = 46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 4 * nlength c + 26 * (nlength c + nlength (x ^ d))\<^sup>2 + (2 * nlength (x ^ d))" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def proof (tform tps: tps5_def tps4_def jk time: assms) show "tps5 = tps4[j + 2 := (\0\\<^sub>N, 1)]" unfolding tps5_def tps4_def using jk tps0 list_update_id[of tps0 "Suc (Suc j)"] by (simp add: list_update_swap) qed lemma tm5': assumes "ttt = 46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2" shows "transforms tm5 tps0 ttt tps5" proof - let ?t = "46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 4 * nlength c + 26 * (nlength c + nlength (x ^ d))\<^sup>2 + (2 * nlength (x ^ d))" have "?t \ 46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 4 * nlength c + 28 * (nlength c + nlength (x ^ d))\<^sup>2" proof - have "2 * nlength (x ^ d) \ 2 * (nlength c + nlength (x ^ d))\<^sup>2" by (meson add_leE eq_imp_le mult_le_mono2 power2_nat_le_imp_le) then show ?thesis by simp qed also have "... \ 46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2" proof - have "4 * nlength c \ 4 * (nlength c + nlength (x ^ d))\<^sup>2" by (simp add: power2_nat_le_eq_le power2_nat_le_imp_le) then show ?thesis by simp qed also have "... = ttt" using assms(1) by simp finally have "?t \ ttt" . then show ?thesis using assms transforms_monotone tm5 by blast qed end (* context x k *) end (* locale *) lemma transforms_tm_monomialI [transforms_intros]: fixes ttt x k :: nat and tps tps' :: "tape list" and j :: tapeidx assumes "j > 0" and "j + 3 < k" and "k = length tps" assumes "tps ! j = (\x\\<^sub>N, 1)" "tps ! (j + 1) = (\0\\<^sub>N, 1)" "tps ! (j + 2) = (\0\\<^sub>N, 1)" "tps ! (j + 3) = (\0\\<^sub>N, 1)" assumes "ttt = 46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2" assumes "tps' = tps[j + 3 := (\c * x ^ d\\<^sub>N, 1)]" shows "transforms (tm_monomial c d j) tps ttt tps'" proof - interpret loc: turing_machine_monomial c d j . show ?thesis using loc.tm5_eq_tm_monomial loc.tm5' loc.tps5_def assms by simp qed subsection \Polynomials\label{s:tm-arithmetic-poly}\ text \ A polynomial is a sum of monomials. In this section we construct for every polynomial function $p$ a Turing machine that on input $x\in\nat$ outputs $p(x)$. According to our definition of polynomials (see Section~\ref{s:tm-basic-bigoh}), we can represent each polynomial by a list of coefficients. The value of such a polynomial with coefficient list @{term cs} on input $x$ is given by the next function. In the following definition, the coefficients of the polynomial are in reverse order, which simplifies the Turing machine later. \ definition polyvalue :: "nat list \ nat \ nat" where "polyvalue cs x \ (\i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0.. polyvalue cs x" using polyvalue_Cons by simp lemma polyvalue_Cons_ge2: "polyvalue (c # cs) x \ c * x ^ (length cs)" using polyvalue_Cons by simp lemma sum_list_const: "(\_\ns. c) = c * length ns" using sum_list_triv[of c ns] by simp lemma polyvalue_le: "polyvalue cs x \ Max (set cs) * length cs * Suc x ^ length cs" proof - define cmax where "cmax = Max (set (rev cs))" have "polyvalue cs x = (\i\[0.. (\i\[0.. cmax" if "i < length cs" for i using that cmax_def by (metis List.finite_set Max_ge length_rev nth_mem) then show ?thesis by (metis (no_types, lifting) atLeastLessThan_iff mult_le_mono1 set_upt sum_list_mono) qed also have "... = cmax * (\i\[0.. cmax * (\i\[0.. cmax * (\i\[0.. Suc x ^ length cs" if "i < length cs" for i using that by (simp add: dual_order.strict_implies_order pow_mono) then show ?thesis by (metis atLeastLessThan_iff mult_le_mono2 set_upt sum_list_mono) qed also have "... = cmax * length cs * Suc x ^ length cs" using sum_list_const[of _ "[0.. cmax * length cs * Suc x ^ length cs" . moreover have "cmax = Max (set cs)" using cmax_def by simp ultimately show ?thesis by simp qed lemma nlength_polyvalue: "nlength (polyvalue cs x) \ nlength (Max (set cs)) + nlength (length cs) + Suc (length cs * nlength (Suc x))" proof - have "nlength (polyvalue cs x) \ nlength (Max (set cs) * length cs * Suc x ^ length cs)" using polyvalue_le nlength_mono by simp also have "... \ nlength (Max (set cs) * length cs) + nlength (Suc x ^ length cs)" using nlength_prod by simp also have "... \ nlength (Max (set cs)) + nlength(length cs) + Suc (length cs * nlength (Suc x))" by (meson add_mono nlength_pow nlength_prod) finally show ?thesis . qed text \ The following Turing machines compute polynomials given as lists of coefficients. If the polynomial is given by coefficients @{term cs}, the TM @{term "tm_polycoef cs j"} expect a number $n$ on tape $j$ and writes $p(n)$ to tape $j + 4$. The tapes $j+1$, $j+2$, and $j + 3$ are auxiliary tapes for use by @{const tm_monomial}. \ fun tm_polycoef :: "nat list \ tapeidx \ machine" where "tm_polycoef [] j = []" | "tm_polycoef (c # cs) j = tm_polycoef cs j ;; (tm_monomial c (length cs) j ;; tm_add (j + 3) (j + 4) ;; tm_setn (j + 3) 0)" lemma tm_polycoef_tm: assumes "k \ 2" and "G \ 4" and "j + 4 < k" and "0 < j" shows "turing_machine k G (tm_polycoef cs j)" proof (induction cs) case Nil then show ?case by (simp add: assms(1) assms(2) turing_machine_def) next case (Cons c cs) moreover have "turing_machine k G (tm_monomial c (length cs) j ;; tm_add (j + 3) (j + 4) ;; tm_setn (j + 3) 0)" using tm_monomial_tm tm_add_tm tm_setn_tm assms by simp ultimately show ?case by simp qed locale turing_machine_polycoef = fixes j :: tapeidx begin definition "tm1 c cs \ tm_monomial c (length cs) j" definition "tm2 c cs \ tm1 c cs ;; tm_add (j + 3) (j + 4)" definition "tm3 c cs \ tm2 c cs ;; tm_setn (j + 3) 0" fun tm4 :: "nat list \ machine" where "tm4 [] = []" | "tm4 (c # cs) = tm4 cs ;; tm3 c cs" lemma tm4_eq_tm_polycoef: "tm4 zs = tm_polycoef zs j" proof (induction zs) case Nil then show ?case by simp next case (Cons z zs) then show ?case by (simp add: tm1_def tm2_def tm3_def) qed context fixes x y k :: nat and tps0 :: "tape list" fixes c :: nat and cs :: "nat list" assumes jk: "0 < j" "j + 4 < k" "k = length tps0" assumes tps0: "tps0 ! j = (\x\\<^sub>N, 1)" "tps0 ! (j + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j + 3) = (\0\\<^sub>N, 1)" "tps0 ! (j + 4) = (\y\\<^sub>N, 1)" begin abbreviation "d \ length cs" definition "tps1 \ tps0 [j + 3 := (\c * x ^ (length cs)\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2" shows "transforms (tm1 c cs) tps0 ttt tps1" unfolding tm1_def by (tform tps: assms jk tps0 tps1_def) definition "tps2 = tps0 [j + 3 := (\c * x ^ (length cs)\\<^sub>N, 1), j + 4 := (\c * x ^ (length cs) + y\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 46 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2 + (3 * max (nlength (c * x ^ d)) (nlength y) + 10)" shows "transforms (tm2 c cs) tps0 ttt tps2" unfolding tm2_def by (tform tps: tps1_def tps2_def jk tps0 time: assms) definition "tps3 \ tps0 [j + 4 := (\c * x ^ d + y\\<^sub>N, 1)]" lemma tm3: assumes "ttt = 66 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2 + 3 * max (nlength (c * x ^ d)) (nlength y) + 2 * nlength (c * x ^ d)" shows "transforms (tm3 c cs) tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps2_def tps3_def jk tps0 time: assms) show "tps3 = tps2[j + 3 := (\0\\<^sub>N, 1)]" using tps3_def tps2_def jk tps0 - by (smt (z3) One_nat_def add_2_eq_Suc add_left_cancel lessI less_numeral_extra(4) list_update_id + by (smt (verit) One_nat_def add_2_eq_Suc add_left_cancel lessI less_numeral_extra(4) list_update_id list_update_overwrite list_update_swap numeral_3_eq_3 numeral_Bit0 plus_1_eq_Suc) qed definition "tps3' \ tps0 [j + 4 := (\c * x ^ length cs + y\\<^sub>N, 1)]" lemma tm3': assumes "ttt = 66 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2 + 5 * max (nlength (c * x ^ d)) (nlength y)" shows "transforms (tm3 c cs) tps0 ttt tps3'" proof - have "66 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2 + 3 * max (nlength (c * x ^ d)) (nlength y) + 2 * nlength (c * x ^ d) \ 66 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2 + 3 * max (nlength (c * x ^ d)) (nlength y) + 2 * max (nlength (c * x ^ d)) (nlength y)" by simp also have "... = 66 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2 + 5 * max (nlength (c * x ^ d)) (nlength y)" by simp finally have "66 + 71 * d + 99 * d ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d))\<^sup>2 + 3 * max (nlength (c * x ^ d)) (nlength y) + 2 * nlength (c * x ^ d) \ ttt" using assms(1) by simp moreover have "tps3' = tps3" using tps3'_def tps3_def by simp ultimately show ?thesis using tm3 transforms_monotone by simp qed end (* context x y k c cs tps0 *) lemma tm3'' [transforms_intros]: fixes c :: nat and cs :: "nat list" fixes x k :: nat and tps0 tps' :: "tape list" assumes "k = length tps0" and "j + 4 < k" and "0 < j" assumes "tps0 ! j = (\x\\<^sub>N, 1)" "tps0 ! (j + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j + 3) = (\0\\<^sub>N, 1)" "tps0 ! (j + 4) = (\polyvalue cs x\\<^sub>N, 1)" assumes "ttt = 66 + 71 * (length cs) + 99 * (length cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ (length cs)))\<^sup>2 + 5 * max (nlength (c * x ^ (length cs))) (nlength (polyvalue cs x))" assumes "tps' = tps0 [j + 4 := (\polyvalue (c # cs) x\\<^sub>N, 1)]" shows "transforms (tm3 c cs) tps0 ttt tps'" using assms tm3'[where ?y="polyvalue cs x"] tps3'_def polyvalue_Cons by simp lemma pow_le_pow_Suc: fixes a b :: nat shows "a ^ b \ Suc a ^ Suc b" proof - have "a ^ b \ Suc a ^ b" by (simp add: power_mono) then show ?thesis by simp qed lemma tm4: fixes x k :: nat and tps0 :: "tape list" fixes cs :: "nat list" assumes "k = length tps0" and "j + 4 < k" and "0 < j" assumes "tps0 ! j = (\x\\<^sub>N, 1)" "tps0 ! (j + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j + 3) = (\0\\<^sub>N, 1)" "tps0 ! (j + 4) = (\0\\<^sub>N, 1)" assumes ttt: "ttt = length cs * (66 + 71 * (length cs) + 99 * (length cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 + 5 * nlength (polyvalue cs x))" shows "transforms (tm4 cs) tps0 ttt (tps0[j + 4 := (\polyvalue cs x\\<^sub>N, 1)])" using ttt proof (induction cs arbitrary: ttt) case Nil then show ?case using polyvalue_Nil transforms_Nil assms by (metis list.size(3) list_update_id mult_is_0 tm4.simps(1)) next case (Cons c cs) note Cons.IH [transforms_intros] have tm4def: "tm4 (c # cs) = tm4 cs ;; tm3 c cs" by simp let ?t1 = "d cs * (66 + 71 * d cs + 99 * d cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (nlength ` set cs) + nlength (Suc x ^ d cs))\<^sup>2 + 5 * nlength (polyvalue cs x))" let ?t2 = "66 + 71 * d cs + 99 * d cs ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d cs))\<^sup>2 + 5 * max (nlength (c * x ^ d cs)) (nlength (polyvalue cs x))" define t where "t = ?t1 + ?t2" have tm4: "transforms (tm4 (c # cs)) tps0 t (tps0[j + 4 := (\polyvalue (c # cs) x\\<^sub>N, 1)])" unfolding tm4def by (tform tps: assms t_def) have "?t1 \ d cs * (66 + 71 * d (c#cs) + 99 * d cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (nlength ` set cs) + nlength (Suc x ^d cs))\<^sup>2 + 5 * nlength (polyvalue cs x))" by simp also have "... \ d cs * (66 + 71 * d (c#cs) + 99 * d (c#cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (nlength ` set cs) + nlength (Suc x ^d cs))\<^sup>2 + 5 * nlength (polyvalue cs x))" by simp also have "... \ d cs * (66 + 71 * d (c#cs) + 99 * d (c#cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (nlength ` set (c#cs)) + nlength (Suc x ^d cs))\<^sup>2 + 5 * nlength (polyvalue cs x))" by simp also have "... \ d cs * (66 + 71 * d (c#cs) + 99 * d (c#cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (nlength ` set (c#cs)) + nlength (Suc x ^d (c#cs)))\<^sup>2 + 5 * nlength (polyvalue cs x))" using nlength_mono by simp also have "... \ d cs * (66 + 71 * d (c#cs) + 99 * d (c#cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (nlength ` set (c#cs)) + nlength (Suc x ^d (c#cs)))\<^sup>2 + 5 * nlength (polyvalue (c#cs) x))" using nlength_mono polyvalue_Cons_ge by simp finally have t1: "?t1 \ d cs * (66 + 71 * d (c#cs) + 99 * d (c#cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (nlength ` set (c#cs)) + nlength (Suc x ^d (c#cs)))\<^sup>2 + 5 * nlength (polyvalue (c#cs) x))" (is "?t1 \ d cs * ?t3") . have "?t2 \ 66 + 71 * d (c # cs) + 99 * d cs ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d cs))\<^sup>2 + 5 * max (nlength (c * x ^ d cs)) (nlength (polyvalue cs x))" by simp also have "... \ 66 + 71 * d (c # cs) + 99 * d (c # cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (nlength c + nlength (x ^ d cs))\<^sup>2 + 5 * max (nlength (c * x ^ d cs)) (nlength (polyvalue cs x))" by simp also have "... \ 66 + 71 * d (c # cs) + 99 * d (c # cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength (c # cs))) + nlength (x ^ d cs))\<^sup>2 + 5 * max (nlength (c * x ^ d cs)) (nlength (polyvalue cs x))" by simp also have "... \ 66 + 71 * d (c # cs) + 99 * d (c # cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength (c # cs))) + nlength (Suc x ^ d (c#cs)))\<^sup>2 + 5 * max (nlength (c * x ^ d cs)) (nlength (polyvalue cs x))" using nlength_mono pow_le_pow_Suc by simp also have "... \ 66 + 71 * d (c # cs) + 99 * d (c # cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength (c # cs))) + nlength (Suc x ^ d (c#cs)))\<^sup>2 + 5 * max (nlength (c * x ^ d cs)) (nlength (polyvalue (c#cs) x))" proof - have "nlength (polyvalue cs x) \ nlength (polyvalue (c#cs) x)" using polyvalue_Cons by (simp add: nlength_mono) then show ?thesis by simp qed also have "... \ 66 + 71 * d (c # cs) + 99 * d (c # cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength (c # cs))) + nlength (Suc x ^ d (c#cs)))\<^sup>2 + 5 * max (nlength (polyvalue (c#cs) x)) (nlength (polyvalue (c#cs) x))" using nlength_mono polyvalue_Cons_ge2 by simp also have "... \ 66 + 71 * d (c # cs) + 99 * d (c # cs) ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength (c # cs))) + nlength (Suc x ^ d (c#cs)))\<^sup>2 + 5 * nlength (polyvalue (c#cs) x)" by simp finally have t2: "?t2 \ ?t3" by simp have "t \ d cs * ?t3 + ?t3" using t1 t2 t_def add_le_mono by blast then have "t \ d (c#cs) * ?t3" by simp moreover have "ttt = d (c#cs) * ?t3" using Cons by simp ultimately have "t \ ttt" by simp then show ?case using tm4 transforms_monotone by simp qed end (* locale turing_machine_polycoef *) text \ The time bound in the previous lemma for @{const tm_polycoef} is a bit unwieldy. It depends not only on the length of the input $x$ but also on the list of coefficients of the polynomial $p$ and on the value $p(x)$. Next we bound this time bound by a simpler expression of the form $d + d\cdot|x|^2$ where $d$ depends only on the polynomial. This is accomplished by the next three lemmas. \ lemma tm_polycoef_time_1: "\d. \x. nlength (polyvalue cs x) \ d + d * nlength x" proof - { fix x have "nlength (polyvalue cs x) \ nlength (Max (set cs)) + nlength (length cs) + Suc (length cs * nlength (Suc x))" using nlength_polyvalue by simp also have "... = nlength (Max (set cs)) + nlength (length cs) + 1 + length cs * nlength (Suc x)" (is "_ = ?a + length cs * nlength (Suc x)") by simp also have "... \ ?a + length cs * (Suc (nlength x))" using nlength_Suc by (meson add_mono_thms_linordered_semiring(2) mult_le_mono2) also have "... = ?a + length cs + length cs * nlength x" (is "_ = ?b + length cs * nlength x") by simp also have "... \ ?b + ?b * nlength x" by (meson add_left_mono le_add2 mult_le_mono1) finally have "nlength (polyvalue cs x) \ ?b + ?b * nlength x" . } then show ?thesis by blast qed lemma tm_polycoef_time_2: "\d. \x. (Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 \ d + d * nlength x ^ 2" proof - { fix x have "(Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 \ (Max (set (map nlength cs)) + Suc (nlength (Suc x) * length cs))\<^sup>2" using nlength_pow by (simp add: mult.commute) also have "... = (Suc (Max (set (map nlength cs))) + nlength (Suc x) * length cs)\<^sup>2" (is "_ = (?a + ?b)^2") by simp also have "... = ?a ^ 2 + 2 * ?a * ?b + ?b ^ 2" by algebra also have "... \ ?a ^ 2 + 2 * ?a * ?b ^ 2 + ?b ^ 2" by (meson add_le_mono dual_order.eq_iff mult_le_mono2 power2_nat_le_imp_le) also have "... \ ?a ^ 2 + (2 * ?a + 1) * ?b ^ 2" by simp also have "... = ?a ^ 2 + (2 * ?a + 1) * (length cs) ^ 2 * nlength (Suc x) ^ 2" by algebra also have "... \ ?a ^ 2 + (2 * ?a + 1) * (length cs) ^ 2 * Suc (nlength x) ^ 2" using nlength_Suc by simp also have "... = ?a ^ 2 + (2 * ?a + 1) * (length cs) ^ 2 * (nlength x ^ 2 + 2 * nlength x + 1)" - by (smt (z3) Suc_eq_plus1 add.assoc mult_2 nat_1_add_1 one_power2 plus_1_eq_Suc power2_sum) + by (smt (verit) Suc_eq_plus1 add.assoc mult_2 nat_1_add_1 one_power2 plus_1_eq_Suc power2_sum) also have "... \ ?a ^ 2 + (2 * ?a + 1) * (length cs) ^ 2 * (nlength x ^ 2 + 2 * nlength x ^ 2 + 1)" proof - have "nlength x ^ 2 + 2 * nlength x + 1 \ nlength x ^ 2 + 2 * nlength x ^ 2 + 1" by (metis add_le_mono1 add_mono_thms_linordered_semiring(2) le_square mult.commute mult_le_mono1 numerals(1) power_add_numeral power_one_right semiring_norm(2)) then show ?thesis by simp qed also have "... = ?a ^ 2 + (2 * ?a + 1) * (length cs) ^ 2 * (3 * nlength x ^ 2 + 1)" by simp also have "... = ?a ^ 2 + (2 * ?a + 1) * (length cs) ^ 2 + (2 * ?a + 1) * (length cs) ^ 2 * 3 * nlength x ^ 2" (is "_ = _ + ?c * nlength x ^ 2") by simp also have "... \ ?a ^ 2 + ?c + ?c * nlength x ^ 2" (is "_ \ ?d + ?c * nlength x ^ 2") by simp also have "... \ ?d + ?d * nlength x ^ 2" by simp finally have "(Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 \ ?d + ?d * nlength x ^ 2" . } then show ?thesis by auto qed lemma tm_polycoef_time_3: "\d. \x. length cs * (66 + 71 * length cs + 99 * length cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 + 5 * nlength (polyvalue cs x)) \ d + d * nlength x ^ 2" proof - obtain d1 where d1: "\x. nlength (polyvalue cs x) \ d1 + d1 * nlength x" using tm_polycoef_time_1 by auto obtain d2 where d2: "\x. (Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 \ d2 + d2 * nlength x ^ 2" using tm_polycoef_time_2 by auto { fix x let ?lhs = " length cs * (66 + 71 * length cs + 99 * length cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 + 5 * nlength (polyvalue cs x))" let ?n = "nlength x" have "?lhs \ length cs * (66 + 71 * length cs + 99 * length cs ^ 3 * ?n ^ 2 + 32 * (d2 + d2 * ?n ^ 2) + 5 * (d1 + d1 * ?n))" using d1 d2 add_le_mono mult_le_mono2 nat_add_left_cancel_le by presburger also have "... \ length cs * (66 + 71 * length cs + 99 * length cs ^ 3 * ?n ^ 2 + 32 * (d2 + d2 * ?n ^ 2) + 5 * (d1 + d1 * ?n ^ 2))" by (simp add: le_square power2_eq_square) also have "... = length cs * (66 + 71 * length cs + 99 * length cs ^ 3 * ?n ^ 2 + 32 * d2 + 32 * d2 * ?n ^ 2 + 5 * d1 + 5 * d1 * ?n ^ 2)" by simp also have "... = length cs * (66 + 71 * length cs + 32 * d2 + 5 * d1 + (99 * length cs ^ 3 + 32 * d2 + 5 * d1) * ?n ^ 2)" by algebra also have "... = length cs * (66 + 71 * length cs + 32 * d2 + 5 * d1) + length cs * (99 * length cs ^ 3 + 32 * d2 + 5 * d1) * ?n ^ 2" (is "_ = ?a + ?b * ?n ^ 2") by algebra also have "... \ max ?a ?b + max ?a ?b * ?n ^ 2" by (simp add: add_mono_thms_linordered_semiring(1)) finally have "?lhs \ max ?a ?b + max ?a ?b * ?n ^ 2" . } then show ?thesis by auto qed text \ According to our definition of @{const polynomial} (see Section~\ref{s:tm-basic-bigoh}) every polynomial has a list of coefficients. Therefore the next definition is well-defined for polynomials $p$. \ definition coefficients :: "(nat \ nat) \ nat list" where "coefficients p \ SOME cs. \n. p n = (\i\[0.. The $d$ in our upper bound of the form $d + d\cdot|x|^2$ for the running time of @{const tm_polycoef} depends on the polynomial. It is given by the next function: \ definition d_polynomial :: "(nat \ nat) \ nat" where "d_polynomial p \ (let cs = rev (coefficients p) in SOME d. \x. length cs * (66 + 71 * length cs + 99 * length cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength cs)) + nlength (Suc x ^ length cs))\<^sup>2 + 5 * nlength (polyvalue cs x)) \ d + d * nlength x ^ 2)" text \ The Turing machine @{const tm_polycoef} has the coefficients of a polynomial as parameter. Next we devise a similar Turing machine that has the polynomial, as a function $\nat \to \nat$, as parameter. \ definition tm_polynomial :: "(nat \ nat) \ tapeidx \ machine" where "tm_polynomial p j \ tm_polycoef (rev (coefficients p)) j" lemma tm_polynomial_tm: assumes "k \ 2" and "G \ 4" and "0 < j" and "j + 4 < k" shows "turing_machine k G (tm_polynomial p j)" using assms tm_polynomial_def tm_polycoef_tm by simp lemma transforms_tm_polynomialI [transforms_intros]: fixes p :: "nat \ nat" and j :: tapeidx fixes k x :: nat and tps tps' :: "tape list" assumes "0 < j" and "k = length tps" and "j + 4 < k" and "polynomial p" assumes "tps ! j = (\x\\<^sub>N, 1)" "tps ! (j + 1) = (\0\\<^sub>N, 1)" "tps ! (j + 2) = (\0\\<^sub>N, 1)" "tps ! (j + 3) = (\0\\<^sub>N, 1)" "tps ! (j + 4) = (\0\\<^sub>N, 1)" assumes "ttt = d_polynomial p + d_polynomial p * nlength x ^ 2" assumes "tps' = tps [j + 4 := (\p x\\<^sub>N, 1)]" shows "transforms (tm_polynomial p j) tps ttt tps'" proof - let ?P = "\x. \n. p n = (\i\[0..cs. ?P cs" using assms(4) polynomial_def by simp ultimately have "?P cs" using someI_ex[of ?P] by blast then have 1: "polyvalue (rev cs) x = p x" using polyvalue_def by simp let ?cs = "rev cs" have "d_polynomial p = (SOME d. \x. length ?cs * (66 + 71 * length ?cs + 99 * length ?cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength ?cs)) + nlength (Suc x ^ length ?cs))\<^sup>2 + 5 * nlength (polyvalue ?cs x)) \ d + d * nlength x ^ 2)" using cs_def coefficients_def d_polynomial_def by simp then have *: "\x. length ?cs * (66 + 71 * length ?cs + 99 * length ?cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength ?cs)) + nlength (Suc x ^ length ?cs))\<^sup>2 + 5 * nlength (polyvalue ?cs x)) \ (d_polynomial p) + (d_polynomial p) * nlength x ^ 2" using tm_polycoef_time_3 someI_ex[OF tm_polycoef_time_3] by presburger let ?ttt = "length ?cs * (66 + 71 * length ?cs + 99 * length ?cs ^ 3 * (nlength x)\<^sup>2 + 32 * (Max (set (map nlength ?cs)) + nlength (Suc x ^ length ?cs))\<^sup>2 + 5 * nlength (polyvalue ?cs x))" interpret loc: turing_machine_polycoef j . have "transforms (loc.tm4 ?cs) tps ?ttt (tps[j + 4 := (\polyvalue ?cs x\\<^sub>N, 1)])" using loc.tm4 assms * by blast then have "transforms (loc.tm4 ?cs) tps ?ttt (tps[j + 4 := (\p x\\<^sub>N, 1)])" using 1 by simp then have "transforms (loc.tm4 ?cs) tps ?ttt tps'" using assms(11) by simp moreover have "loc.tm4 ?cs = tm_polynomial p j" using tm_polynomial_def loc.tm4_eq_tm_polycoef coefficients_def cs_def by simp ultimately have "transforms (tm_polynomial p j) tps ?ttt tps'" by simp then show "transforms (tm_polynomial p j) tps ttt tps'" using * assms(10) transforms_monotone by simp qed subsection \Division by two\ text \ In order to divide a number by two, a Turing machine can shift all symbols on the tape containing the number to the left, of course without overwriting the start symbol. The next command implements the left shift. It scans the tape $j$ from right to left and memorizes the current symbol on the last tape. It works very similar to @{const cmd_double} only in the opposite direction. Upon reaching the start symbol, it moves the head one cell to the right. \ definition cmd_halve :: "tapeidx \ command" where "cmd_halve j rs \ (if rs ! j = 1 then 1 else 0, (map (\i. if i = j then if rs ! j = \ then (rs ! i, Right) else if last rs = \ then (\, Left) else (tosym (todigit (last rs)), Left) else if i = length rs - 1 then (tosym (todigit (rs ! j)), Stay) else (rs ! i, Stay)) [0.. 4" and "0 < j" and "j < k" shows "turing_command (Suc k) 1 G (cmd_halve j)" proof show "\gs. length gs = Suc k \ length ([!!] cmd_halve j gs) = length gs" using cmd_halve_def by simp moreover have "0 \ Suc k - 1" using assms by simp ultimately show "\gs. length gs = Suc k \ 0 < Suc k \ cmd_halve j gs [.] 0 = gs ! 0" using assms cmd_halve_def by (smt (verit) One_nat_def ab_semigroup_add_class.add_ac(1) diff_Suc_1 length_map neq0_conv nth_map nth_upt plus_1_eq_Suc prod.sel(1) prod.sel(2)) show "cmd_halve j gs [.] j' < G" if "length gs = Suc k" "(\i. i < length gs \ gs ! i < G)" "j' < length gs" for gs j' proof - have "cmd_halve j gs [!] j' = (if j' = j then if gs ! j = \ then (gs ! j', Right) else if last gs = \ then (\, Left) else (tosym (todigit (last gs)), Left) else if j' = length gs - 1 then (tosym (todigit (gs ! j)), Stay) else (gs ! j', Stay))" using cmd_halve_def that(3) by simp moreover consider "j' = j" | "j' = k" | "j' \ j \ j' \ k" by auto ultimately show ?thesis using that assms by (cases) simp_all qed show "\gs. length gs = Suc k \ [*] (cmd_halve j gs) \ 1" using cmd_halve_def by simp qed lemma sem_cmd_halve_2: assumes "j < k" and "bit_symbols xs" and "length tps = Suc k" and "i \ length xs" and "i > 0" and "z = \ \ z = \" and "tps ! j = (\xs\, i)" and "tps ! k = \z\" and "tps' = tps[j := tps ! j |:=| z |-| 1, k := \xs ! (i - 1)\]" shows "sem (cmd_halve j) (0, tps) = (0, tps')" proof (rule semI) show "proper_command (Suc k) (cmd_halve j)" using cmd_halve_def by simp show "length tps = Suc k" "length tps' = Suc k" using assms(3,9) by simp_all define rs where "rs = read tps" then have lenrs: "length rs = Suc k" using assms(3) read_length by simp have rsj: "rs ! j = xs ! (i - 1)" using rs_def assms tapes_at_read' contents_inbounds by (metis fst_conv le_imp_less_Suc less_imp_le_nat snd_conv) then have rsj': "rs ! j > 1" using assms Suc_1 Suc_diff_1 Suc_le_lessD by (metis eval_nat_numeral(3) less_Suc_eq) then show "fst (cmd_halve j (read tps)) = 0" using cmd_halve_def rs_def by simp have lastrs: "last rs = z" using assms rs_def onesie_read tapes_at_read' by (metis diff_Suc_1 last_conv_nth length_0_conv lenrs lessI nat.simps(3)) show "act (cmd_halve j (read tps) [!] j') (tps ! j') = tps' ! j'" if "j' < Suc k" for j' proof - have "j' < length rs" using that lenrs by simp then have *: "cmd_halve j rs [!] j' = (if j' = j then if rs ! j = \ then (rs ! j', Right) else if last rs = \ then (\, Left) else (tosym (todigit (last rs)), Left) else if j' = length rs - 1 then (tosym (todigit (rs ! j)), Stay) else (rs ! j', Stay))" using cmd_halve_def by simp consider "j' = j" | "j' = k" | "j' \ j \ j' \ k" by auto then show ?thesis proof (cases) case 1 then have "cmd_halve j (read tps) [!] j' = (tosym (todigit (last rs)), Left)" using rs_def rsj' lastrs * assms(6) by auto then have "cmd_halve j (read tps) [!] j' = (z, Left)" using lastrs assms(6) by auto moreover have "tps' ! j' = tps ! j |:=| z |-| 1" using 1 assms(1,3,9) by simp ultimately show ?thesis using act_Left' 1 that rs_def by metis next case 2 then have "cmd_halve j (read tps) [!] j' = (tosym (todigit (rs ! j)), Stay)" using rs_def * lenrs assms(1) by simp moreover have "tps' ! j' = \xs ! (i - 1)\" using assms 2 by simp moreover have "tps ! j' = \z\" using assms 2 by simp moreover have "tosym (todigit (rs ! j)) = xs ! (i - 1)" proof - have "xs ! (i - 1) = \ \ xs ! (i - 1) = \" using rsj rs_def assms by simp then show ?thesis using One_nat_def add_2_eq_Suc' numeral_3_eq_3 rsj by presburger qed ultimately show ?thesis using act_onesie by simp next case 3 then show ?thesis using * act_Stay that assms lenrs rs_def by simp qed qed qed lemma sem_cmd_halve_1: assumes "j < k" and "bit_symbols xs" and "length tps = Suc k" and "0 < length xs" and "tps ! j = (\xs\, length xs)" and "tps ! k = \\\" and "tps' = tps[j := tps ! j |:=| \ |-| 1, k := \xs ! (length xs - 1)\]" shows "sem (cmd_halve j) (0, tps) = (0, tps')" proof (rule semI) show "proper_command (Suc k) (cmd_halve j)" using cmd_halve_def by simp show "length tps = Suc k" "length tps' = Suc k" using assms(3,7) by simp_all define rs where "rs = read tps" then have lenrs: "length rs = Suc k" using assms(3) read_length by simp have rsj: "rs ! j = xs ! (length xs - 1)" using rs_def assms tapes_at_read' contents_inbounds by (metis One_nat_def fst_conv le_eq_less_or_eq le_imp_less_Suc snd_conv) then have rsj': "rs ! j > 1" using assms(2,4) by (metis One_nat_def Suc_1 diff_less lessI less_add_Suc2 numeral_3_eq_3 plus_1_eq_Suc) then show "fst (cmd_halve j (read tps)) = \" using cmd_halve_def rs_def by simp have lastrs: "last rs = \" using assms rs_def onesie_read tapes_at_read' by (metis diff_Suc_1 last_conv_nth length_0_conv lenrs lessI nat.simps(3)) show "act (cmd_halve j (read tps) [!] j') (tps ! j') = tps' ! j'" if "j' < Suc k" for j' proof - have "j' < length rs" using that lenrs by simp then have *: "cmd_halve j rs [!] j' = (if j' = j then if rs ! j = \ then (rs ! j', Right) else if last rs = \ then (\, Left) else (tosym (todigit (last rs)), Left) else if j' = length rs - 1 then (tosym (todigit (rs ! j)), Stay) else (rs ! j', Stay))" using cmd_halve_def by simp consider "j' = j" | "j' = k" | "j' \ j \ j' \ k" by auto then show ?thesis proof (cases) case 1 then have "cmd_halve j (read tps) [!] j' = (\, Left)" using rs_def rsj' lastrs * by simp then show ?thesis using act_Left' 1 that rs_def assms(1,3,7) by simp next case 2 then have "cmd_halve j (read tps) [!] j' = (tosym (todigit (rs ! j)), Stay)" using rs_def * lenrs assms(1) by simp moreover have "tps' ! j' = \xs ! (length xs - 1)\" using assms 2 by simp moreover have "tps ! j' = \\\" using assms 2 by simp ultimately show ?thesis using act_onesie assms 2 that rs_def rsj - by (smt (z3) One_nat_def Suc_1 add_2_eq_Suc' diff_less numeral_3_eq_3 zero_less_one) + by (smt (verit) One_nat_def Suc_1 add_2_eq_Suc' diff_less numeral_3_eq_3 zero_less_one) next case 3 then show ?thesis using * act_Stay that assms lenrs rs_def by simp qed qed qed lemma sem_cmd_halve_0: assumes "j < k" and "length tps = Suc k" and "tps ! j = (\xs\, 0)" and "tps ! k = \z\" and "tps' = tps[j := tps ! j |+| 1, k := \\\]" shows "sem (cmd_halve j) (0, tps) = (1, tps')" proof (rule semI) show "proper_command (Suc k) (cmd_halve j)" using cmd_halve_def by simp show "length tps = Suc k" "length tps' = Suc k" using assms(2,5) by simp_all show "fst (cmd_halve j (read tps)) = 1" using cmd_halve_def assms contents_at_0 tapes_at_read' by (smt (verit) fst_conv le_eq_less_or_eq not_less not_less_eq snd_conv) show "act (cmd_halve j (read tps) [!] j') (tps ! j') = tps' ! j'" if "j' < Suc k" for j' proof - define gs where "gs = read tps" then have "length gs = Suc k" using assms by (simp add: read_length) then have "j' < length gs" using that by simp then have *: "cmd_halve j gs [!] j' = (if j' = j then if gs ! j = \ then (gs ! j', Right) else if last gs = \ then (\, Left) else (tosym (todigit (last gs)), Left) else if j' = length gs - 1 then (tosym (todigit (gs ! j)), Stay) else (gs ! j', Stay))" using cmd_halve_def by simp have gsj: "gs ! j = \" using gs_def assms(1,2,3) by (metis contents_at_0 fstI less_Suc_eq sndI tapes_at_read') consider "j' = j" | "j' = k" | "j' \ j \ j' \ k" by auto then show ?thesis proof (cases) case 1 then have "cmd_halve j (read tps) [!] j' = (gs ! j', Right)" using gs_def gsj * by simp then show ?thesis using act_Right assms 1 that gs_def by (metis length_list_update lessI nat_neq_iff nth_list_update) next case 2 then have "cmd_halve j (read tps) [!] j' = (tosym (todigit (gs ! j)), Stay)" using gs_def * \length gs = Suc k\ assms(1) by simp moreover have "tps' ! j' = \\\" using assms 2 by simp moreover have "tps ! j' = \z\" using assms 2 by simp ultimately show ?thesis using act_onesie assms 2 that gs_def gsj by (smt (verit, best) One_nat_def Suc_1 add_2_eq_Suc' less_Suc_eq_0_disj less_numeral_extra(3) nat.inject numeral_3_eq_3) next case 3 then show ?thesis using * act_Stay that assms(2,5) \length gs = Suc k\ gs_def by simp qed qed qed definition tm_halve :: "tapeidx \ machine" where "tm_halve j \ [cmd_halve j]" lemma tm_halve_tm: assumes "k \ 2" and "G \ 4" and "0 < j" and "j < k" shows "turing_machine (Suc k) G (tm_halve j)" using tm_halve_def turing_command_halve assms by auto lemma exe_cmd_halve_0: assumes "j < k" and "length tps = Suc k" and "tps ! j = (\xs\, 0)" and "tps ! k = \z\" and "tps' = tps[j := tps ! j |+| 1, k := \\\]" shows "exe (tm_halve j) (0, tps) = (1, tps')" using assms sem_cmd_halve_0 tm_halve_def exe_lt_length by simp lemma execute_cmd_halve_0: assumes "j < k" and "length tps = Suc k" and "tps ! j = (\[]\, 0)" and "tps ! k = \\\" and "tps' = tps[j := tps ! j |+| 1, k := \\\]" shows "execute (tm_halve j) (0, tps) 1 = (1, tps')" using tm_halve_def exe_lt_length sem_cmd_halve_0 assms by simp definition shift :: "tape \ nat \ tape" where "shift tp y \ (\x. if x \ y then (fst tp) x else (fst tp) (Suc x), y)" lemma shift_update: "y > 0 \ shift tp y |:=| (fst tp) (Suc y) |-| 1 = shift tp (y - 1)" unfolding shift_def by fastforce lemma shift_contents_0: assumes "length xs > 0" shows "shift (\xs\, length xs) 0 = (\tl xs\, 0)" proof - have "shift (\xs\, length xs) 0 = (\drop 1 xs\, 0)" using shift_def contents_def by fastforce then show ?thesis by (simp add: drop_Suc) qed lemma proper_bit_symbols: "bit_symbols ws \ proper_symbols ws" by auto lemma bit_symbols_shift: assumes "t < length ws" and "bit_symbols ws" shows "|.| (shift (\ws\, length ws) (length ws - t)) \ 1" using assms shift_def contents_def nat_neq_iff proper_bit_symbols by simp lemma exe_cmd_halve_1: assumes "j < k" and "length tps = Suc k" and "bit_symbols xs" and "length xs > 0" and "tps ! j = (\xs\, length xs)" and "tps ! k = \\\" and "tps' = tps[j := tps ! j |:=| \ |-| 1, k := \xs ! (length xs - 1)\]" shows "exe (tm_halve j) (0, tps) = (0, tps')" using tm_halve_def exe_lt_length sem_cmd_halve_1 assms by simp lemma shift_contents_eq_take_drop: assumes "length xs > 0" and "ys = take i xs @ drop (Suc i) xs" and "i > 0" and "i < length xs" shows "shift (\xs\, length xs) i = (\ys\, i)" proof - have "shift (\xs\, length xs) i = (\x. if x \ i then \xs\ x else \xs\ (Suc x), i)" using shift_def by auto moreover have "(\x. if x \ i then \xs\ x else \xs\ (Suc x)) = \take i xs @ drop (Suc i) xs\" (is "?l = ?r") proof fix x consider "x = 0" | "0 < x \ x \ i" | "i < x \ x \ length xs - 1" | "length xs - 1 < x" by linarith then show "?l x = ?r x" proof (cases) case 1 then show ?thesis using assms contents_def by simp next case 2 then have "?l x = \xs\ x" by simp then have lhs: "?l x = xs ! (x - 1)" using assms 2 by simp have "?r x = (take i xs @ drop (Suc i) xs) ! (x - 1)" using assms 2 by auto then have "?r x = xs ! (x - 1)" using assms(4) 2 by (metis diff_less le_eq_less_or_eq length_take less_trans min_absorb2 nth_append nth_take zero_less_one) then show ?thesis using lhs by simp next case 3 then have "?l x = \xs\ (Suc x)" by simp then have lhs: "?l x = xs ! x" using 3 assms by auto have "?r x = (take i xs @ drop (Suc i) xs) ! (x - 1)" using assms 3 by auto then have "?r x = drop (Suc i) xs ! (x - 1 - i)" using assms(3,4) 3 - by (smt (z3) Suc_diff_1 dual_order.strict_trans length_take less_Suc_eq min_absorb2 nat_less_le nth_append) + by (smt (verit) Suc_diff_1 dual_order.strict_trans length_take less_Suc_eq min_absorb2 nat_less_le nth_append) then have "?r x = xs ! x" using assms 3 by simp then show ?thesis using lhs by simp next case 4 then show ?thesis using contents_def by auto qed qed ultimately show ?thesis using assms(2) by simp qed lemma exe_cmd_halve_2: assumes "j < k" and "bit_symbols xs" and "length tps = Suc k" and "i \ length xs" and "i > 0" and "z = \ \ z = \" and "tps ! j = (\xs\, i)" and "tps ! k = \z\" and "tps' = tps[j := tps ! j |:=| z |-| 1, k := \xs ! (i - 1)\]" shows "exe (tm_halve j) (0, tps) = (0, tps')" using tm_halve_def exe_lt_length sem_cmd_halve_2 assms by simp lemma shift_contents_length_minus_1: assumes "length xs > 0" shows "shift (\xs\, length xs) (length xs - 1) = (\xs\, length xs) |:=| \ |-| 1" using contents_def shift_def assms by fastforce lemma execute_tm_halve_1_less: assumes "j < k" and "length tps = Suc k" and "bit_symbols xs" and "length xs > 0" and "tps ! j = (\xs\, length xs)" and "tps ! k = \\\" and "t \ 1" and "t \ length xs" shows "execute (tm_halve j) (0, tps) t = (0, tps [j := shift (tps ! j) (length xs - t), k := \xs ! (length xs - t)\])" using assms(7,8) proof (induction t rule: nat_induct_at_least) case base have "execute (tm_halve j) (0, tps) 1 = exe (tm_halve j) (0, tps)" by simp also have "... = (0, tps[j := tps ! j |:=| \ |-| 1, k := \xs ! (length xs - 1)\])" using assms exe_cmd_halve_1 by simp also have "... = (0, tps[j := shift (tps ! j) (length xs - 1), k := \xs ! (length xs - 1)\])" using shift_contents_length_minus_1 assms(4,5) by simp finally show ?case . next case (Suc t) then have "t < length xs" by simp let ?ys = "take (length xs - t) xs @ drop (Suc (length xs - t)) xs" have "execute (tm_halve j) (0, tps) (Suc t) = exe (tm_halve j) (execute (tm_halve j) (0, tps) t)" by simp also have "... = exe (tm_halve j) (0, tps [j := shift (tps ! j) (length xs - t), k := \xs ! (length xs - t)\])" using Suc by simp also have "... = exe (tm_halve j) (0, tps [j := shift (\xs\, length xs) (length xs - t), k := \xs ! (length xs - t)\])" using assms(5) by simp also have "... = exe (tm_halve j) (0, tps [j := (\?ys\, length xs - t), k := \xs ! (length xs - t)\])" (is "_ = exe _ (0, ?tps)") using shift_contents_eq_take_drop Suc assms by simp also have "... = (0, ?tps [j := ?tps ! j |:=| (xs ! (length xs - t)) |-| 1, k := \?ys ! (length xs - t - 1)\])" proof - let ?i = "length xs - t" let ?z = "xs ! ?i" have 1: "bit_symbols ?ys" using assms(3) by (intro bit_symbols_append) simp_all have 2: "length ?tps = Suc k" using assms(2) by simp have 3: "?i \ length ?ys" using Suc assms by simp have 4: "?i > 0" using Suc assms by simp have 5: "?z = 2 \ ?z = 3" using assms(3,4) Suc by simp have 6: "?tps ! j = (\?ys\, ?i)" using assms(1,2) by simp have 7: "?tps ! k = \?z\" using assms(2) by simp then show ?thesis using exe_cmd_halve_2[OF assms(1) 1 2 3 4 5 6 7] by simp qed also have "... = (0, tps [j := ?tps ! j |:=| (xs ! (length xs - t)) |-| 1, k := \?ys ! (length xs - t - 1)\])" - using assms by (smt (z3) list_update_overwrite list_update_swap) + using assms by (smt (verit) list_update_overwrite list_update_swap) also have "... = (0, tps [j := (\?ys\, length xs - t) |:=| (xs ! (length xs - t)) |-| 1, k := \?ys ! (length xs - t - 1)\])" using assms(1,2) by simp also have "... = (0, tps [j := shift (\xs\, length xs) (length xs - Suc t), k := \xs ! (length xs - (Suc t))\])" proof - have "(\?ys\, length xs - t) |:=| xs ! (length xs - t) |-| 1 = shift (\xs\, length xs) (length xs - t) |:=| (xs ! (length xs - t)) |-| 1" using shift_contents_eq_take_drop One_nat_def Suc Suc_le_lessD \t < length xs\ assms(4) diff_less zero_less_diff by presburger also have "... = shift (\xs\, length xs) (length xs - Suc t)" using shift_update[of "length xs - t" "(\xs\, length xs)"] assms Suc by simp finally have "(\?ys\, length xs - t) |:=| xs ! (length xs - t) |-| 1 = shift (\xs\, length xs) (length xs - Suc t)" . moreover have "?ys ! (length xs - t - 1) = xs ! (length xs - Suc t)" using Suc assms \t < length xs\ by (metis (no_types, lifting) diff_Suc_eq_diff_pred diff_Suc_less diff_commute diff_less length_take min_less_iff_conj nth_append nth_take zero_less_diff zero_less_one) ultimately show ?thesis by simp qed also have "... = (0, tps [j := shift (tps ! j) (length xs - (Suc t)), k := \xs ! (length xs - (Suc t))\])" using assms(5) by simp finally show ?case . qed lemma execute_tm_halve_1: assumes "j < k" and "length tps = Suc k" and "bit_symbols xs" and "length xs > 0" and "tps ! j = (\xs\, length xs)" and "tps ! k = \\\" and "tps' = tps[j := (\tl xs\, 1), k := \\\]" shows "execute (tm_halve j) (0, tps) (Suc (length xs)) = (1, tps')" proof - have "execute (tm_halve j) (0, tps) (length xs) = (0, tps[j := shift (tps ! j) 0, k := \xs ! 0\])" using execute_tm_halve_1_less[OF assms(1-6), where ?t="length xs"] assms(4) by simp also have "... = (0, tps[j := shift (\xs\, length xs) 0, k := \xs ! 0\])" using assms(5) by simp also have "... = (0, tps[j := (\tl xs\, 0), k := \xs ! 0\])" using shift_contents_0 assms(4) by simp finally have "execute (tm_halve j) (0, tps) (length xs) = (0, tps[j := (\tl xs\, 0), k := \xs ! 0\])" . then have "execute (tm_halve j) (0, tps) (Suc (length xs)) = exe (tm_halve j) (0, tps[j := (\tl xs\, 0), k := \xs ! 0\])" (is "_ = exe _ (0, ?tps)") by simp also have "... = (1, ?tps[j := (\tl xs\, 0) |+| 1, k := \\\])" using assms(1,2) exe_cmd_halve_0 by simp also have "... = (1, tps[j := (\tl xs\, 0) |+| 1, k := \\\])" using assms(1,2) by (metis (no_types, opaque_lifting) list_update_overwrite list_update_swap) also have "... = (1, tps[j := (\tl xs\, 1), k := \\\])" by simp finally show ?thesis using assms(7) by simp qed lemma execute_tm_halve: assumes "j < k" and "length tps = Suc k" and "bit_symbols xs" and "tps ! j = (\xs\, length xs)" and "tps ! k = \\\" and "tps' = tps[j := (\tl xs\, 1), k := \\\]" shows "execute (tm_halve j) (0, tps) (Suc (length xs)) = (1, tps')" using execute_cmd_halve_0 execute_tm_halve_1 assms by (cases "length xs = 0") simp_all lemma transforms_tm_halve: assumes "j < k" and "length tps = Suc k" and "bit_symbols xs" and "tps ! j = (\xs\, length xs)" and "tps ! k = \\\" and "tps' = tps[j := (\tl xs\, 1), k := \\\]" shows "transforms (tm_halve j) tps (Suc (length xs)) tps'" using execute_tm_halve assms tm_halve_def transforms_def transits_def by auto lemma transforms_tm_halve2: assumes "j < k" and "length tps = k" and "bit_symbols xs" and "tps ! j = (\xs\, length xs)" and "tps' = tps[j := (\tl xs\, 1)]" shows "transforms (tm_halve j) (tps @ [\\\]) (Suc (length xs)) (tps' @ [\\\])" proof - let ?tps = "tps @ [\\\]" let ?tps' = "tps' @ [\\\]" have "?tps ! j = (\xs\, length xs)" "?tps ! k = \\\" using assms by (simp_all add: nth_append) moreover have "?tps' ! j = (\tl xs\, 1)" "?tps' ! k = \\\" using assms by (simp_all add: nth_append) moreover have "length ?tps = Suc k" using assms(2) by simp ultimately show ?thesis using assms transforms_tm_halve[OF assms(1), where ?tps="?tps" and ?tps'="?tps'" and ?xs=xs] by (metis length_list_update list_update_append1 list_update_length) qed text \ The next Turing machine removes the memorization tape from @{const tm_halve}. \ definition tm_halve' :: "tapeidx \ machine" where "tm_halve' j \ cartesian (tm_halve j) 4" lemma bounded_write_tm_halve: assumes "j < k" shows "bounded_write (tm_halve j) k 4" unfolding bounded_write_def proof standard+ fix q :: nat and rs :: "symbol list" assume q: "q < length (tm_halve j)" and lenrs: "length rs = Suc k" have "k < length rs" using lenrs by simp then have "cmd_halve j rs [!] k = (if k = j then if rs ! j = \ then (rs ! k, Right) else if last rs = \ then (\, Left) else (tosym (todigit (last rs)), Left) else if k = length rs - 1 then (tosym (todigit (rs ! j)), Stay) else (rs ! k, Stay))" using cmd_halve_def by simp then have "cmd_halve j rs [!] k = (tosym (todigit (rs ! j)), Stay)" using assms lenrs by simp then have "cmd_halve j rs [.] k = tosym (todigit (rs ! j))" by simp moreover have "(tm_halve j ! q) rs [.] k = cmd_halve j rs [.] k" using tm_halve_def q by simp ultimately show "(tm_halve j ! q) rs [.] k < 4" by simp qed lemma immobile_tm_halve: assumes "j < k" shows "immobile (tm_halve j) k (Suc k)" proof standard+ fix q :: nat and rs :: "symbol list" assume q: "q < length (tm_halve j)" and lenrs: "length rs = Suc k" have "k < length rs" using lenrs by simp then have "cmd_halve j rs [!] k = (if k = j then if rs ! j = \ then (rs ! k, Right) else if last rs = \ then (\, Left) else (tosym (todigit (last rs)), Left) else if k = length rs - 1 then (tosym (todigit (rs ! j)), Stay) else (rs ! k, Stay))" using cmd_halve_def by simp then have "cmd_halve j rs [!] k = (tosym (todigit (rs ! j)), Stay)" using assms lenrs by simp then have "cmd_halve j rs [~] k = Stay" by simp moreover have "(tm_halve j ! q) rs [~] k = cmd_halve j rs [~] k" using tm_halve_def q by simp ultimately show "(tm_halve j ! q) rs [~] k = Stay" by simp qed lemma tm_halve'_tm: assumes "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_halve' j)" using tm_halve'_def tm_halve_tm assms cartesian_tm by simp lemma transforms_tm_halve' [transforms_intros]: assumes "j > 0" and "j < k" and "length tps = k" and "bit_symbols xs" and "tps ! j = (\xs\, length xs)" and "tps' = tps[j := (\tl xs\, 1)]" shows "transforms (tm_halve' j) tps (Suc (length xs)) tps'" unfolding tm_halve'_def proof (rule cartesian_transforms_onesie[OF tm_halve_tm immobile_tm_halve _ _ bounded_write_tm_halve assms(3), where ?G=4]; (simp add: assms)?) show "2 \ k" and "2 \ k" using assms(1,2) by simp_all show "transforms (tm_halve j) (tps @ [\Suc 0\]) (Suc (length xs)) (tps[j := (\tl xs\, Suc 0)] @ [\\\])" using transforms_tm_halve2 assms by simp qed lemma num_tl_div_2: "num (tl xs) = num xs div 2" proof (cases "xs = []") case True then show ?thesis by (simp add: num_def) next case False then have *: "xs = hd xs # tl xs" by simp then have "num xs = todigit (hd xs) + 2 * num (tl xs)" using num_Cons by metis then show ?thesis by simp qed lemma canrepr_div_2: "canrepr (n div 2) = tl (canrepr n)" using canreprI canrepr canonical_canrepr num_tl_div_2 canonical_tl by (metis hd_Cons_tl list.sel(2)) corollary nlength_times2: "nlength (2 * n) \ Suc (nlength n)" using canrepr_div_2[of "2 * n"] by simp corollary nlength_times2plus1: "nlength (2 * n + 1) \ Suc (nlength n)" using canrepr_div_2[of "2 * n + 1"] by simp text \ The next Turing machine is the one we actually use to divide a number by two. First it moves to the end of the symbol sequence representing the number, then it applies @{const tm_halve'}. \ definition tm_div2 :: "tapeidx \ machine" where "tm_div2 j \ tm_right_until j {\} ;; tm_left j ;; tm_halve' j" lemma tm_div2_tm: assumes "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_div2 j)" unfolding tm_div2_def using tm_right_until_tm tm_left_tm tm_halve'_tm assms by simp locale turing_machine_div2 = fixes j :: tapeidx begin definition "tm1 \ tm_right_until j {\}" definition "tm2 \ tm1 ;; tm_left j" definition "tm3 \ tm2 ;; tm_halve' j" lemma tm3_eq_tm_div2: "tm3 = tm_div2 j" unfolding tm3_def tm2_def tm1_def tm_div2_def by simp context fixes tps0 :: "tape list" and k n :: nat assumes jk: "0 < j" "j < k" "length tps0 = k" and tps0: "tps0 ! j = (\n\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j := (\n\\<^sub>N, Suc (nlength n))]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (nlength n)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps1_def jk tps0 time: assms) have "rneigh (\n\\<^sub>N, Suc 0) {\} (nlength n)" proof (intro rneighI) show "fst (\n\\<^sub>N, Suc 0) (snd (\n\\<^sub>N, Suc 0) + nlength n) \ {\}" using contents_def by simp show "\n'. n' < nlength n \ fst (\n\\<^sub>N, Suc 0) (snd (\n\\<^sub>N, Suc 0) + n') \ {\}" using bit_symbols_canrepr contents_def contents_outofbounds proper_symbols_canrepr by (metis One_nat_def Suc_leI add_diff_cancel_left' fst_eqD less_Suc_eq_0_disj less_nat_zero_code plus_1_eq_Suc singletonD snd_conv) qed then show "rneigh (tps0 ! j) {\} (nlength n)" using tps0 by simp qed definition "tps2 \ tps0 [j := (\n\\<^sub>N, nlength n)]" lemma tm2 [transforms_intros]: assumes "ttt = 2 + nlength n" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: tps1_def tps2_def jk assms) definition "tps3 \ tps0 [j := (\n div 2\\<^sub>N, 1)]" lemma tm3: assumes "ttt = 2 * nlength n + 3" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps3_def tps2_def tps0 jk time: assms) show "bit_symbols (canrepr n)" using bit_symbols_canrepr . show "tps3 = tps2[j := (\tl (canrepr n)\, 1)]" using tps3_def tps2_def jk tps0 canrepr_div_2 by simp qed end end (* locale turing_machine_div2 *) lemma transforms_tm_div2I [transforms_intros]: fixes tps tps' :: "tape list" and ttt k n :: nat and j :: tapeidx assumes "0 < j" "j < k" and "length tps = k" and "tps ! j = (\n\\<^sub>N, 1)" assumes "ttt = 2 * nlength n + 3" assumes "tps' = tps[j := (\n div 2\\<^sub>N, 1)]" shows "transforms (tm_div2 j) tps ttt tps'" proof - interpret loc: turing_machine_div2 j . show ?thesis using loc.tm3_eq_tm_div2 loc.tm3 loc.tps3_def assms by simp qed subsection \Modulo two\ text \ In this section we construct a Turing machine that writes to tape $j_2$ the symbol @{text \} or @{text \} depending on whether the number on tape $j_1$ is odd or even. If initially tape $j_2$ contained at most one symbol, it will contain the numbers~1 or~0. \ lemma canrepr_odd: "odd n \ canrepr n ! 0 = \" proof - assume "odd n" then have "0 < n" by presburger then have len: "length (canrepr n) > 0" using nlength_0 by simp then have "canrepr n ! 0 = \ \ canrepr n ! 0 = \" using bit_symbols_canrepr by fastforce then show "canrepr n ! 0 = \" using prepend_2_even len canrepr `odd n` `0 < n` by (metis gr0_implies_Suc length_Suc_conv nth_Cons_0) qed lemma canrepr_even: "even n \ 0 < n \ canrepr n ! 0 = \" proof - assume "even n" "0 < n" then have len: "length (canrepr n) > 0" using nlength_0 by simp then have "canrepr n ! 0 = \ \ canrepr n ! 0 = \" using bit_symbols_canrepr by fastforce then show "canrepr n ! 0 = \" using prepend_3_odd len canrepr `even n` `0 < n` by (metis gr0_implies_Suc length_Suc_conv nth_Cons_0) qed definition "tm_mod2 j1 j2 \ tm_trans2 j1 j2 (\z. if z = \ then \ else \)" lemma tm_mod2_tm: assumes "k \ 2" and "G \ 4" and "0 < j2" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_mod2 j1 j2)" unfolding tm_mod2_def using assms tm_trans2_tm by simp lemma transforms_tm_mod2I [transforms_intros]: assumes "j1 < length tps" and "0 < j2" and "j2 < length tps" and "b \ 1" assumes "tps ! j1 = (\n\\<^sub>N, 1)" and "tps ! j2 = (\b\\<^sub>N, 1)" assumes "tps' = tps[j2 := (\n mod 2\\<^sub>N, 1)]" shows "transforms (tm_mod2 j1 j2) tps 1 tps'" proof - let ?f = "\z::symbol. if z = \ then \ else \" let ?tps = "tps[j2 := tps ! j2 |:=| (?f (tps :.: j1))]" have *: "transforms (tm_mod2 j1 j2) tps 1 ?tps" using transforms_tm_trans2I assms tm_mod2_def by metis have "tps :.: j1 = \" if "odd n" using that canrepr_odd assms(5) contents_def by (metis One_nat_def diff_Suc_1 fst_conv gr_implies_not0 ncontents_1_blank_iff_zero odd_pos snd_conv) moreover have "tps :.: j1 = \" if "even n" and "n > 0" using that canrepr_even assms(5) contents_def by (metis One_nat_def diff_Suc_1 fst_conv gr_implies_not0 ncontents_1_blank_iff_zero snd_conv) moreover have "tps :.: j1 = \" if "n = 0" using that canrepr_even assms(5) contents_def by simp ultimately have "tps :.: j1 = \ \ odd n" by linarith then have f: "?f (tps :.: j1) = \ \ odd n" by simp have tps_j2: "tps ! j2 |:=| (?f (tps :.: j1)) = ((\b\\<^sub>N)(1 := (?f (tps :.: j1))), 1)" using assms by simp have "tps ! j2 |:=| (?f (tps :.: j1)) = (\n mod 2\\<^sub>N, 1)" proof (cases "even n") case True then have "tps ! j2 |:=| (?f (tps :.: j1)) = ((\b\\<^sub>N)(1 := 0), 1)" using f tps_j2 by auto also have "... = (\[]\, 1)" proof (cases "b = 0") case True then have "\b\\<^sub>N = \[]\" using canrepr_0 by simp then show ?thesis by auto next case False then have "\b\\<^sub>N = \[\]\" using canrepr_1 assms(4) by (metis One_nat_def bot_nat_0.extremum_uniqueI le_Suc_eq) then show ?thesis by (metis One_nat_def append.simps(1) append_Nil2 contents_append_update contents_blank_0 list.size(3)) qed also have "... = (\0\\<^sub>N, 1)" using canrepr_0 by simp finally show ?thesis using True by auto next case False then have "tps ! j2 |:=| (?f (tps :.: j1)) = ((\b\\<^sub>N)(1 := \), 1)" using f tps_j2 by auto also have "... = (\[\]\, 1)" proof (cases "b = 0") case True then have "\b\\<^sub>N = \[]\" using canrepr_0 by simp then show ?thesis by (metis One_nat_def append.simps(1) contents_snoc list.size(3)) next case False then have "\b\\<^sub>N = \[\]\" using canrepr_1 assms(4) by (metis One_nat_def bot_nat_0.extremum_uniqueI le_Suc_eq) then show ?thesis by auto qed also have "... = (\1\\<^sub>N, 1)" using canrepr_1 by simp also have "... = (\n mod 2\\<^sub>N, 1)" using False by (simp add: mod2_eq_if) finally show ?thesis by auto qed then show ?thesis using * assms(7) by auto qed subsection \Boolean operations\ text \ In order to support Boolean operations, we represent the value True by the number~1 and False by~0. \ abbreviation bcontents :: "bool \ (nat \ symbol)" ("\_\\<^sub>B") where "\b\\<^sub>B \ \if b then 1 else 0\\<^sub>N" text \ A tape containing a number contains the number~0 iff.\ there is a blank in cell number~1. \ lemma read_ncontents_eq_0: assumes "tps ! j = (\n\\<^sub>N, 1)" and "j < length tps" shows "(read tps) ! j = \ \ n = 0" using assms tapes_at_read'[of j tps] ncontents_1_blank_iff_zero by (metis prod.sel(1) prod.sel(2)) subsubsection \And\ text \ The next Turing machine, when given two numbers $a, b \in \{0, 1\}$ on tapes $j_1$ and $j_2$, writes to tape $j_1$ the number~1 if $a = b = 1$; otherwise it writes the number~0. In other words, it overwrites tape $j_1$ with the logical AND of the two tapes. \ definition tm_and :: "tapeidx \ tapeidx \ machine" where "tm_and j1 j2 \ IF \rs. rs ! j1 = \ \ rs ! j2 = \ THEN tm_write j1 \ ELSE [] ENDIF" lemma tm_and_tm: assumes "k \ 2" and "G \ 4" and "0 < j1" and "j1 < k" shows "turing_machine k G (tm_and j1 j2)" using tm_and_def tm_write_tm Nil_tm assms turing_machine_branch_turing_machine by simp locale turing_machine_and = fixes j1 j2 :: tapeidx begin context fixes tps0 :: "tape list" and k :: nat and a b :: nat assumes ab: "a < 2" "b < 2" assumes jk: "j1 < k" "j2 < k" "j1 \ j2" "0 < j1" "length tps0 = k" assumes tps0: "tps0 ! j1 = (\a\\<^sub>N, 1)" "tps0 ! j2 = (\b\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j1 := (\a = 1 \ b = 1\\<^sub>B, 1)]" lemma tm: "transforms (tm_and j1 j2) tps0 3 tps1" unfolding tm_and_def proof (tform) have "read tps0 ! j1 = \canrepr a\ 1" using jk tps0 tapes_at_read'[of j1 tps0] by simp then have 1: "read tps0 ! j1 = \ \ a = 1" using ab canrepr_odd contents_def ncontents_1_blank_iff_zero by (metis (mono_tags, lifting) One_nat_def diff_Suc_1 less_2_cases_iff odd_one) have "read tps0 ! j2 = \canrepr b\ 1" using jk tps0 tapes_at_read'[of j2 tps0] by simp then have 2: "read tps0 ! j2 = \ \ b = 1" using ab canrepr_odd contents_def ncontents_1_blank_iff_zero by (metis (mono_tags, lifting) One_nat_def diff_Suc_1 less_2_cases_iff odd_one) show "tps1 = tps0" if "\ (read tps0 ! j1 = \ \ read tps0 ! j2 = \)" proof - have "a = (if a = 1 \ b = 1 then 1 else 0)" using that 1 2 ab jk by (metis One_nat_def less_2_cases_iff read_ncontents_eq_0 tps0(2)) then have "tps0 ! j1 = (\a = 1 \ b = 1\\<^sub>B, 1)" using tps0 by simp then show ?thesis unfolding tps1_def using list_update_id[of tps0 j1] by simp qed show "tps1 = tps0[j1 := tps0 ! j1 |:=| \]" if "read tps0 ! j1 = \ \ read tps0 ! j2 = \" proof - have "(if a = 1 \ b = 1 then 1 else 0) = 0" using that 1 2 by simp moreover have "tps0 ! j1 |:=| \ = (\0\\<^sub>N, 1)" proof (cases "a = 0") case True then show ?thesis using tps0 jk by auto next case False then have "a = 1" using ab by simp then have "\a\\<^sub>N = \[\]\" using canrepr_1 by simp moreover have "(\[\]\, 1) |:=| \ = (\[]\, 1)" using contents_def by auto ultimately have "(\a\\<^sub>N, 1) |:=| \ = (\0\\<^sub>N, 1)" using ncontents_0 by presburger then show ?thesis using tps0 jk by simp qed ultimately have "tps0 ! j1 |:=| \ = (\a = 1 \ b = 1\\<^sub>B, 1)" by (smt (verit, best)) then show ?thesis unfolding tps1_def by auto qed qed end (* context *) end (* locale *) lemma transforms_tm_andI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps :: "tape list" and k :: nat and a b :: nat assumes "a < 2" "b < 2" assumes "length tps = k" assumes "j1 < k" "j2 < k" "j1 \ j2" "0 < j1" assumes "tps ! j1 = (\a\\<^sub>N, 1)" "tps ! j2 = (\b\\<^sub>N, 1)" assumes "tps' = tps [j1 := (\a = 1 \ b = 1\\<^sub>B, 1)]" shows "transforms (tm_and j1 j2) tps 3 tps'" proof - interpret loc: turing_machine_and j1 j2 . show ?thesis using assms loc.tps1_def loc.tm by simp qed subsubsection \Not\ text \ The next Turing machine turns the number~1 into~0 and vice versa. \ definition tm_not :: "tapeidx \ machine" where "tm_not j \ IF \rs. rs ! j = \ THEN tm_write j \ ELSE tm_write j \ ENDIF" lemma tm_not_tm: assumes "k \ 2" and "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_not j)" using tm_not_def tm_write_tm assms turing_machine_branch_turing_machine by simp locale turing_machine_not = fixes j :: tapeidx begin context fixes tps0 :: "tape list" and k :: nat and a :: nat assumes a: "a < 2" assumes jk: "j < k" "length tps0 = k" assumes tps0: "tps0 ! j = (\a\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j := (\a \ 1\\<^sub>B, 1)]" lemma tm: "transforms (tm_not j) tps0 3 tps1" unfolding tm_not_def proof (tform) have *: "read tps0 ! j = \ \ a = 0" using read_ncontents_eq_0 jk tps0 by simp show "tps1 = tps0[j := tps0 ! j |:=| \]" if "read tps0 ! j = \" proof - have "a = 0" using a that * by simp then have "(\if a = 1 then 0 else 1\\<^sub>N, 1) = (\1\\<^sub>N, 1)" by simp moreover have "tps0 ! j |:=| \ = (\1\\<^sub>N, 1)" using tps0 canrepr_0 canrepr_1 `a = 0` contents_snoc by (metis One_nat_def append.simps(1) fst_conv list.size(3) snd_conv) ultimately have "tps0[j := tps0 ! j |:=| \] = tps0[j := (\a \ 1\\<^sub>B, 1)]" by auto then show ?thesis using tps1_def by simp qed show "tps1 = tps0[j := tps0 ! j |:=| \]" if "read tps0 ! j \ \" proof - have "a = 1" using a that * by simp then have "(\if a = 1 then 0 else 1\\<^sub>N, 1) = (\0\\<^sub>N, 1)" by simp moreover have "tps0 ! j |:=| \ = (\0\\<^sub>N, 1)" using tps0 canrepr_0 canrepr_1 `a = 1` contents_snoc by (metis Suc_1 append_self_conv2 contents_blank_0 fst_eqD fun_upd_upd nat.inject nlength_0_simp numeral_2_eq_2 snd_eqD) ultimately have "tps0[j := tps0 ! j |:=| \] = tps0[j := (\a \ 1\\<^sub>B, 1)]" by auto then show ?thesis using tps1_def by simp qed qed end (* context *) end (* locale *) lemma transforms_tm_notI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and k :: nat and a :: nat assumes "j < k" "length tps = k" and "a < 2" assumes "tps ! j = (\a\\<^sub>N, 1)" assumes "tps' = tps [j := (\a \ 1\\<^sub>B, 1)]" shows "transforms (tm_not j) tps 3 tps'" proof - interpret loc: turing_machine_not j . show ?thesis using assms loc.tps1_def loc.tm by simp qed end \ No newline at end of file diff --git a/thys/Cook_Levin/Aux_TM_Reducing.thy b/thys/Cook_Levin/Aux_TM_Reducing.thy --- a/thys/Cook_Levin/Aux_TM_Reducing.thy +++ b/thys/Cook_Levin/Aux_TM_Reducing.thy @@ -1,4253 +1,4254 @@ chapter \Auxiliary Turing machines for reducing $\NP$ languages to \SAT{}\label{s:Aux_TM}\ theory Aux_TM_Reducing imports Reducing begin text \ In the previous chapter we have seen how to reduce a language $L\in\NP$ to \SAT{} by constructing for every string $x$ a CNF formula $\Phi$ that is satisfiable iff.\ $x\in L$. To complete the Cook-Levin theorem it remains to show that there is a polynomial-time Turing machine that on input $x$ outputs $\Phi$. Constructing such a TM will be the subject of the rest of this article and conclude in the next chapter. This chapter introduces several TMs used in the construction. \ section \Generating literals\ text \ Our representation of CNF formulas as lists of lists of numbers is based on a representation of literals as numbers. Our function @{const literal_n} encodes the positive literal $v_i$ as the number $2i + 1$ and the negative literal $\bar v_i$ as $2i$. We already have the Turing machine @{const tm_times2} to cover the second case. Now we build a TM for the first case, that is, for doubling and incrementing. \null \ definition tm_times2incr :: "tapeidx \ machine" where "tm_times2incr j \ tm_times2 j ;; tm_incr j" lemma tm_times2incr_tm: assumes "0 < j" and "j < k" and "G \ 4" shows "turing_machine k G (tm_times2incr j)" unfolding tm_times2incr_def using tm_times2_tm tm_incr_tm assms by simp lemma transforms_tm_times2incrI [transforms_intros]: fixes j :: tapeidx fixes k :: nat and tps tps' :: "tape list" assumes "k \ 2" and "j > 0" and "j < k" and "length tps = k" assumes "tps ! j = (\n\\<^sub>N, 1)" assumes "t = 12 + 4 * nlength n" assumes "tps' = tps[j := (\Suc (2 * n)\\<^sub>N, 1)]" shows "transforms (tm_times2incr j) tps t tps'" proof - define tt where "tt = 10 + (2 * nlength n + 2 * nlength (2 * n))" have "transforms (tm_times2incr j) tps tt tps'" unfolding tm_times2incr_def by (tform tps: tt_def assms) moreover have "tt \ t" proof - have "tt = 10 + 2 * nlength n + 2 * nlength (2 * n)" using tt_def by simp also have "... \ 10 + 2 * nlength n + 2 * (Suc (nlength n))" proof - have "nlength (2 * n) \ Suc (nlength n)" by (metis eq_imp_le gr0I le_SucI nat_0_less_mult_iff nlength_even_le) then show ?thesis by simp qed also have "... = 12 + 4 * nlength n" by simp finally show ?thesis using assms(6) by simp qed ultimately show ?thesis using transforms_monotone by simp qed lemma literal_n_rename: assumes "v div 2 < length \" shows "2 * \ ! (v div 2) + v mod 2 = (literal_n \ rename \) (n_literal v)" proof (cases "even v") case True then show ?thesis using n_literal_def assms by simp next case False then show ?thesis using n_literal_def assms by simp presburger qed text \ Combining @{const tm_times2} and @{const tm_times2incr}, the next Turing machine accepts a variable index $i$ on tape $j_1$ and a flag $b$ on tape $j_2$ and outputs on tape $j_1$ the encoding of the positive literal $v_i$ or the negative literal $\bar v_i$ if $b$ is positive or zero, respectively. \ definition tm_to_literal :: "tapeidx \ tapeidx \ machine" where "tm_to_literal j1 j2 \ IF \rs. rs ! j2 = \ THEN tm_times2 j1 ELSE tm_times2incr j1 ENDIF" lemma tm_to_literal_tm: assumes "k \ 2" and "G \ 4" and "0 < j1" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_to_literal j1 j2)" unfolding tm_to_literal_def using assms tm_times2_tm tm_times2incr_tm turing_machine_branch_turing_machine by simp lemma transforms_tm_to_literalI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and ttt k i b :: nat assumes "0 < j1" "j1 < k" "j2 < k" "2 \ k" "length tps = k" assumes "tps ! j1 = (\i\\<^sub>N, 1)" "tps ! j2 = (\b\\<^sub>N, 1)" assumes "ttt = 13 + 4 * nlength i" assumes "tps' = tps [j1 := (\2 * i + (if b = 0 then 0 else 1)\\<^sub>N, 1)]" shows "transforms (tm_to_literal j1 j2) tps ttt tps'" unfolding tm_to_literal_def proof (tform tps: assms read_ncontents_eq_0) show "5 + 2 * nlength i + 2 \ ttt" and "12 + 4 * nlength i + 1 \ ttt" using assms(8) by simp_all qed section \A Turing machine for relabeling formulas\ text \ In order to construct $\Phi_9$, we must construct CNF formulas $\chi_t$, which have the form $\rho(\psi)$ or $\rho'(\psi')$. So we need a Turing machine for relabeling formulas. In this section we devise a Turing machine that gets a substitution $\sigma$ and a CNF formula $\phi$ and outputs $\sigma(\phi)$. In order to bound its running time we first prove some bounds on the length of relabeled formulas. \ subsection \The length of relabeled formulas\ text \ First we bound the length of the representation of a single relabeled clause. In the following lemma the assumption ensures that the substitution $\sigma$ has a value for every variable in the clause. \ lemma nllength_rename: assumes "\v\set clause. v div 2 < length \" shows "nllength (map (literal_n \ rename \) (n_clause clause)) \ length clause * Suc (nllength \)" proof (cases "\ = []") case True then show ?thesis using assms n_clause_def by simp next case False let ?f = "literal_n \ rename \ \ n_literal" have *: "map (literal_n \ rename \) (n_clause clause) = map ?f clause" using n_clause_def by simp have "nlength (2 * n + 1) \ Suc (nlength n)" for n using nlength_times2plus1 by simp then have "nlength (2 * Max (set \) + 1) \ Suc (nlength (Max (set \)))" by simp moreover have "nlength (Max (set \)) \ nllength \ - 1" using False member_le_nllength_1 by simp ultimately have "nlength (2 * Max (set \) + 1) \ Suc (nllength \ - 1)" by simp then have **: "nlength (2 * Max (set \) + 1) \ nllength \" using nllength_gr_0 False by simp have "?f n \ 2 * (\ ! (n div 2)) + 1" if "n div 2 < length \" for n using n_literal_def by (cases "even n") simp_all then have "?f v \ 2 * (\ ! (v div 2)) + 1" if "v \ set clause" for v using assms that by simp moreover have "\ ! (v div 2) \ Max (set \)" if "v \ set clause" for v using that assms by simp ultimately have "?f v \ 2 * Max (set \) + 1" if "v \ set clause" for v using that by fastforce then have "n \ 2 * Max (set \) + 1" if "n \ set (map ?f clause)" for n using that by auto then have "nllength (map ?f clause) \ Suc (nlength (2 * Max (set \) + 1)) * length (map ?f clause)" using nllength_le_len_mult_max by blast also have "... = Suc (nlength (2 * Max (set \) + 1)) * length clause" by simp also have "... \ Suc (nllength \) * length clause" using ** by simp finally have "nllength (map ?f clause) \ Suc (nllength \) * length clause" . then show ?thesis using * by (metis mult.commute) qed text \ Our upper bound for the length of the symbol representation of a relabeled formula is fairly crude. It is basically the length of the string resulting from replacing every symbol of the original formula by the entire substitution. \ lemma nlllength_relabel: assumes "\clause\set \. \v\set (clause_n clause). v div 2 < length \" shows "nlllength (formula_n (relabel \ \)) \ Suc (nllength \) * nlllength (formula_n \)" using assms proof (induction \) case Nil then show ?case by (simp add: relabel_def) next case (Cons clause \) let ?nclause = "clause_n clause" have "\v\set ?nclause. v div 2 < length \" using Cons.prems by simp then have "nllength (map (literal_n \ rename \) (n_clause ?nclause)) \ length ?nclause * Suc (nllength \)" using nllength_rename by simp then have "nllength (map (literal_n \ rename \) clause) \ length clause * Suc (nllength \)" using clause_n_def n_clause_n by simp moreover have "map (literal_n \ rename \) clause = clause_n (map (rename \) clause)" using clause_n_def by simp ultimately have *: "nllength (clause_n (map (rename \) clause)) \ length clause * Suc (nllength \)" by simp have "formula_n (relabel \ (clause # \)) = clause_n (map (rename \) clause) # formula_n (relabel \ \)" by (simp add: formula_n_def relabel_def) then have "nlllength (formula_n (relabel \ (clause # \))) = nllength (clause_n (map (rename \) clause)) + 1 + nlllength (formula_n (relabel \ \))" using nlllength_Cons by simp also have "... \ length clause * Suc (nllength \) + 1 + nlllength (formula_n (relabel \ \))" using * by simp also have "... \ length clause * Suc (nllength \) + 1 + Suc (nllength \) * nlllength (formula_n \)" using Cons by (metis add_mono_thms_linordered_semiring(2) insert_iff list.set(2)) also have "... = 1 + Suc (nllength \) * (length clause + nlllength (formula_n \))" by algebra also have "... \ Suc (nllength \) * (1 + length clause + nlllength (formula_n \))" by simp also have "... \ Suc (nllength \) * (1 + nllength (clause_n clause) + nlllength (formula_n \))" using length_le_nllength n_clause_def n_clause_n by (metis add_Suc_shift add_le_cancel_right length_map mult_le_mono2 plus_1_eq_Suc) also have "... = Suc (nllength \) * (nlllength (formula_n (clause # \)))" using formula_n_def nlllength_Cons by simp finally show ?case . qed text \ A simple sufficient condition for the assumption in the previous lemma. \ lemma variables_\: assumes "variables \ \ {..}" shows "\clause\set \.\v\set (clause_n clause). v div 2 < length \" proof standard+ fix clause :: clause and v :: nat assume clause: "clause \ set \" and v: "v \ set (clause_n clause)" obtain i where i: "i < length clause" "v = literal_n (clause ! i)" using v clause_n_def by (metis in_set_conv_nth length_map nth_map) then have clause_i: "clause ! i = n_literal v" using n_literal_n by simp show "v div 2 < length \" proof (cases "even v") case True then have "clause ! i = Neg (v div 2)" using clause_i n_literal_def by simp then have "\c\set \. Neg (v div 2) \ set c" using clause i(1) by (metis nth_mem) then have "v div 2 \ variables \" using variables_def by auto then show ?thesis using assms by auto next case False then have "clause ! i = Pos (v div 2)" using clause_i n_literal_def by simp then have "\c\set \. Pos (v div 2) \ set c" using clause i(1) by (metis nth_mem) then have "v div 2 \ variables \" using variables_def by auto then show ?thesis using assms by auto qed qed text \ Combining the previous two lemmas yields this upper bound: \ lemma nlllength_relabel_variables: assumes "variables \ \ {..}" shows "nlllength (formula_n (relabel \ \)) \ Suc (nllength \) * nlllength (formula_n \)" using assms variables_\ nlllength_relabel by blast subsection \Relabeling clauses\ text \ Relabeling a CNF formula is accomplished by relabeling each of its clauses. In this section we devise a Turing machine for relabeling clauses. The TM accepts on tape $j$ a list of numbers representing a substitution $\sigma$ and on tape $j + 1$ a clause represented as a list of numbers. It outputs on tape $j + 2$ the relabeled clause, consuming the original clause on tape $j + 1$ in the process. \ definition tm_relabel_clause :: "tapeidx \ machine" where "tm_relabel_clause j \ WHILE [] ; \rs. rs ! (j + 1) \ \ DO tm_nextract \ (j + 1) (j + 3) ;; tm_mod2 (j + 3) (j + 4) ;; tm_div2 (j + 3) ;; tm_nth_inplace j (j + 3) \ ;; tm_to_literal (j + 3) (j + 4) ;; tm_append (j + 2) (j + 3) ;; tm_setn (j + 3) 0 ;; tm_setn (j + 4) 0 DONE ;; tm_cr (j + 2) ;; tm_erase_cr (j + 1)" lemma tm_relabel_clause_tm: assumes "G \ 5" and "j + 4 < k" and "0 < j" shows "turing_machine k G (tm_relabel_clause j)" unfolding tm_relabel_clause_def using assms tm_nextract_tm tm_mod2_tm tm_div2_tm tm_nth_inplace_tm tm_to_literal_tm tm_append_tm tm_setn_tm tm_cr_tm tm_erase_cr_tm Nil_tm turing_machine_loop_turing_machine by simp locale turing_machine_relabel_clause = fixes j :: tapeidx begin definition "tm1 \ tm_nextract \ (j + 1) (j + 3)" definition "tm2 \ tm1 ;; tm_mod2 (j + 3) (j + 4)" definition "tm3 \ tm2 ;; tm_div2 (j + 3)" definition "tm4 \ tm3 ;; tm_nth_inplace j (j + 3) \" definition "tm5 \ tm4 ;; tm_to_literal (j + 3) (j + 4)" definition "tm6 \ tm5 ;; tm_append (j + 2) (j + 3)" definition "tm7 = tm6 ;; tm_setn (j + 3) 0" definition "tm8 \ tm7 ;; tm_setn (j + 4) 0" definition "tmL \ WHILE [] ; \rs. rs ! (j + 1) \ \ DO tm8 DONE" definition "tm9 \ tmL ;; tm_cr (j + 2)" definition "tm10 \ tm9 ;; tm_erase_cr (j + 1)" lemma tm10_eq_tm_relabel_clause: "tm10 = tm_relabel_clause j" unfolding tm_relabel_clause_def tm3_def tmL_def tm5_def tm4_def tm1_def tm2_def tm6_def tm7_def tm8_def tm9_def tm10_def by simp context fixes tps0 :: "tape list" and k :: nat and \ clause :: "nat list" assumes clause: "\v\set clause. v div 2 < length \" assumes jk: "0 < j" "j + 4 < k" "length tps0 = k" assumes tps0: "tps0 ! j = (\\\\<^sub>N\<^sub>L, 1)" "tps0 ! (j + 1) = (\clause\\<^sub>N\<^sub>L, 1)" "tps0 ! (j + 2) = (\[]\\<^sub>N\<^sub>L, 1)" "tps0 ! (j + 3) = (\0\\<^sub>N, 1)" "tps0 ! (j + 4) = (\0\\<^sub>N, 1)" begin definition "tpsL t \ tps0 [j + 1 := nltape' clause t, j + 2 := nltape (take t (map literal_n (map (rename \) (n_clause clause))))]" lemma tpsL_eq_tps0: "tpsL 0 = tps0" using tpsL_def tps0 jk nllength_Nil by (metis One_nat_def list_update_id take0) definition "tps1 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take t (map literal_n (map (rename \) (n_clause clause)))), j + 3 := (\clause ! t\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "t < length clause" and "ttt = 12 + 2 * nlength (clause ! t)" shows "transforms tm1 (tpsL t) ttt (tps1 t)" unfolding tm1_def proof (tform tps: assms(1) tpsL_def tps1_def tps0 jk) show "ttt = 12 + 2 * nlength 0 + 2 * nlength (clause ! t)" using assms(2) by simp qed definition "tps2 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take t (map literal_n (map (rename \) (n_clause clause)))), j + 3 := (\clause ! t\\<^sub>N, 1), j + 4 := (\(clause ! t) mod 2\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "t < length clause" and "ttt = 13 + 2 * nlength (clause ! t)" shows "transforms tm2 (tpsL t) ttt (tps2 t)" unfolding tm2_def by (tform tps: assms tps1_def tps2_def tps0 jk) definition "tps3 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take t (map literal_n (map (rename \) (n_clause clause)))), j + 3 := (\clause ! t div 2\\<^sub>N, 1), j + 4 := (\clause ! t mod 2\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "t < length clause" and "ttt = 16 + 4 * nlength (clause ! t)" shows "transforms tm3 (tpsL t) ttt (tps3 t)" unfolding tm3_def by (tform tps: assms(1) tps2_def tps3_def jk time: assms(2)) definition "tps4 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take t (map literal_n (map (rename \) (n_clause clause)))), j + 3 := (\\ ! (clause ! t div 2)\\<^sub>N, 1), j + 4 := (\clause ! t mod 2\\<^sub>N, 1)]" lemma tm4 [transforms_intros]: assumes "t < length clause" and "ttt = 28 + 4 * nlength (clause ! t) + 18 * (nllength \)\<^sup>2" shows "transforms tm4 (tpsL t) ttt (tps4 t)" unfolding tm4_def proof (tform tps: assms(1) tps0 tps3_def tps4_def jk clause time: assms(2)) show "tps4 t = (tps3 t)[j + 3 := (\\ ! (clause ! t div 2)\\<^sub>N, 1)]" unfolding tps4_def tps3_def by (simp add: list_update_swap[of "j + 3"]) qed definition "tps5 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take t (map literal_n (map (rename \) (n_clause clause)))), j + 3 := (\2 * (\ ! (clause ! t div 2)) + clause ! t mod 2\\<^sub>N, 1), j + 4 := (\clause ! t mod 2\\<^sub>N, 1)]" lemma tm5 [transforms_intros]: assumes "t < length clause" and "ttt = 41 + 4 * nlength (clause ! t) + 18 * (nllength \)\<^sup>2 + 4 * nlength (\ ! (clause ! t div 2))" shows "transforms tm5 (tpsL t) ttt (tps5 t)" unfolding tm5_def by (tform tps: assms(1) tps0 tps4_def tps5_def jk time: assms(2)) definition "tps6 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take (Suc t) (map literal_n (map (rename \) (n_clause clause)))), j + 3 := (\2 * (\ ! (clause ! t div 2)) + clause ! t mod 2\\<^sub>N, 1), j + 4 := (\clause ! t mod 2\\<^sub>N, 1)]" lemma tm6: assumes "t < length clause" and "ttt = 47 + 4 * nlength (clause ! t) + 18 * (nllength \)\<^sup>2 + 4 * nlength (\ ! (clause ! t div 2)) + 2 * nlength (2 * \ ! (clause ! t div 2) + clause ! t mod 2)" shows "transforms tm6 (tpsL t) ttt (tps6 t)" unfolding tm6_def proof (tform tps: assms(1) tps0 tps5_def tps6_def jk) have "2 * \ ! (clause ! t div 2) + clause ! t mod 2 = (literal_n \ rename \) (n_literal (clause ! t))" using clause assms(1) literal_n_rename by simp then have "2 * \ ! (clause ! t div 2) + clause ! t mod 2 = (map (literal_n \ rename \) (n_clause clause)) ! t" using assms(1) by (simp add: n_clause_def) then have "take t (map (literal_n \ rename \) (n_clause clause)) @ [2 * \ ! (clause ! t div 2) + clause ! t mod 2] = take (Suc t) (map (literal_n \ rename \) (n_clause clause))" by (simp add: assms(1) n_clause_def take_Suc_conv_app_nth) then show "tps6 t = (tps5 t) [j + 2 := nltape (take t (map (literal_n \ rename \) (n_clause clause)) @ [2 * \ ! (clause ! t div 2) + clause ! t mod 2])]" unfolding tps5_def tps6_def by (simp only: list_update_overwrite list_update_swap_less[of "j + 2"]) simp show "ttt = 41 + 4 * nlength (clause ! t) + 18 * (nllength \)\<^sup>2 + 4 * nlength (\ ! (clause ! t div 2)) + (7 + nllength (take t (map (literal_n \ rename \) (n_clause clause))) - Suc (nllength (take t (map (literal_n \ rename \) (n_clause clause)))) + 2 * nlength (2 * \ ! (clause ! t div 2) + clause ! t mod 2))" using assms(2) by simp qed lemma nlength_\1: assumes "t < length clause" shows "nlength (clause ! t) \ nllength \" proof - have "clause ! t div 2 < length \" using clause assms(1) by simp then have "nlength (clause ! t div 2) < length \" using nlength_le_n by (meson leD le_less_linear order.trans) then have "nlength (clause ! t) \ length \" using canrepr_div_2 by simp then show "nlength (clause ! t) \ nllength \" using length_le_nllength by (meson dual_order.trans mult_le_mono2) qed lemma nlength_\2: assumes "t < length clause" shows "nlength (\ ! (clause ! t div 2)) \ nllength \" using assms clause member_le_nllength nth_mem by simp lemma nlength_\3: assumes "t < length clause" shows "nlength (2 * \ ! (clause ! t div 2) + clause ! t mod 2) \ Suc (nllength \)" proof - have "nlength (2 * \ ! (clause ! t div 2) + clause ! t mod 2) \ nlength (2 * \ ! (clause ! t div 2) + 1)" using nlength_mono by simp also have "... \ Suc (nlength (\ ! (clause ! t div 2)))" using nlength_times2plus1 by (meson dual_order.trans) finally show ?thesis using nlength_\2 assms by fastforce qed lemma tm6' [transforms_intros]: assumes "t < length clause" and "ttt = 49 + 28 * nllength \ ^ 2" shows "transforms tm6 (tpsL t) ttt (tps6 t)" proof - let ?ttt = "47 + 4 * nlength (clause ! t) + 18 * (nllength \)\<^sup>2 + 4 * nlength (\ ! (clause ! t div 2)) + 2 * nlength (2 * \ ! (clause ! t div 2) + clause ! t mod 2)" have "?ttt \ 47 + 4 * nllength \ + 18 * (nllength \)\<^sup>2 + 4 * nllength \ + 2 * Suc (nllength \)" using nlength_\1 nlength_\3 nlength_\2 assms(1) by fastforce also have "... = 49 + 10 * nllength \ + 18 * (nllength \)\<^sup>2" by simp also have "... \ 49 + 10 * nllength \ ^ 2 + 18 * (nllength \)\<^sup>2" using linear_le_pow by simp also have "... = 49 + 28 * nllength \ ^ 2" by simp finally have "?ttt \ 49 + 28 * nllength \ ^ 2" . then show ?thesis using tm6 assms transforms_monotone by blast qed definition "tps7 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take (Suc t) (map literal_n (map (rename \) (n_clause clause)))), j + 4 := (\clause ! t mod 2\\<^sub>N, 1)]" lemma tm7: assumes "t < length clause" and "ttt = 59 + 28 * nllength \ ^ 2 + 2 * nlength (2 * \ ! (clause ! t div 2) + clause ! t mod 2)" shows "transforms tm7 (tpsL t) ttt (tps7 t)" unfolding tm7_def proof (tform tps: assms(1) tps0 tps6_def tps7_def jk time: assms(2)) show "tps7 t = (tps6 t)[j + 3 := (\0\\<^sub>N, 1)]" using tps6_def tps7_def tps0 jk - by (smt (z3) add_left_cancel list_update_id list_update_overwrite list_update_swap num.simps(8) + by (smt (verit) add_left_cancel list_update_id list_update_overwrite list_update_swap num.simps(8) numeral_eq_iff one_eq_numeral_iff semiring_norm(84)) qed lemma tm7' [transforms_intros]: assumes "t < length clause" and "ttt = 61 + 30 * nllength \ ^ 2" shows "transforms tm7 (tpsL t) ttt (tps7 t)" proof - let ?ttt = "59 + 28 * nllength \ ^ 2 + 2 * nlength (2 * \ ! (clause ! t div 2) + clause ! t mod 2)" have "?ttt \ 59 + 28 * nllength \ ^ 2 + 2 * Suc (nllength \)" using nlength_\3 assms(1) by fastforce also have "... = 61 + 28 * nllength \ ^ 2 + 2 * nllength \" by simp also have "... \ 61 + 30 * nllength \ ^ 2" using linear_le_pow by simp finally have "?ttt \ 61 + 30 * nllength \ ^ 2" . then show ?thesis using assms tm7 transforms_monotone by blast qed definition "tps8 t \ tps0 [j + 1 := nltape' clause (Suc t), j + 2 := nltape (take (Suc t) (map literal_n (map (rename \) (n_clause clause))))]" lemma tm8: assumes "t < length clause" and "ttt = 61 + 30 * (nllength \)\<^sup>2 + (10 + 2 * nlength (clause ! t mod 2))" shows "transforms tm8 (tpsL t) ttt (tps8 t)" unfolding tm8_def proof (tform tps: assms(1) tps0 tps7_def tps8_def jk time: assms(2)) show "tps8 t = (tps7 t)[j + 4 := (\0\\<^sub>N, 1)]" using tps7_def tps8_def tps0 jk - by (smt (z3) add_left_imp_eq list_update_id list_update_overwrite list_update_swap numeral_eq_iff + by (smt (verit) add_left_imp_eq list_update_id list_update_overwrite list_update_swap numeral_eq_iff numeral_eq_one_iff semiring_norm(85) semiring_norm(87)) qed lemma tm8' [transforms_intros]: assumes "t < length clause" and "ttt = 71 + 32 * (nllength \)\<^sup>2" shows "transforms tm8 (tpsL t) ttt (tpsL (Suc t))" proof - have "nlength (clause ! t mod 2) \ nllength \" using assms(1) nlength_\1 by (meson mod_less_eq_dividend nlength_mono order.trans) then have "nlength (clause ! t mod 2) \ nllength \ ^ 2" using linear_le_pow by (metis nat_le_linear power2_nat_le_imp_le verit_la_disequality) then have "61 + 30 * (nllength \)\<^sup>2 + (10 + 2 * nlength (clause ! t mod 2)) \ ttt" using assms(2) by simp then have "transforms tm8 (tpsL t) ttt (tps8 t)" using assms(1) tm8 transforms_monotone by blast moreover have "tps8 t = tpsL (Suc t)" using tps8_def tpsL_def by simp ultimately show ?thesis by simp qed lemma tmL [transforms_intros]: assumes "ttt = length clause * (73 + 32 * (nllength \)\<^sup>2) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length clause))" unfolding tmL_def proof (tform) let ?t = "71 + 32 * (nllength \)\<^sup>2" show "read (tpsL t) ! (j + 1) \ \" if "t < length clause" for t proof - have "tpsL t ! (j + 1) = nltape' clause t" using tpsL_def jk by simp then show ?thesis using nltape'_tape_read that tapes_at_read' tpsL_def jk - by (smt (z3) Suc_eq_plus1 leD length_list_update less_add_same_cancel1 less_trans_Suc zero_less_numeral) + by (smt (verit) Suc_eq_plus1 leD length_list_update less_add_same_cancel1 less_trans_Suc zero_less_numeral) qed show "\ read (tpsL (length clause)) ! (j + 1) \ \" proof - have "tpsL (length clause) ! (j + 1) = nltape' clause (length clause)" using tpsL_def jk by simp then show ?thesis using nltape'_tape_read tapes_at_read' tpsL_def jk - by (smt (z3) Suc_eq_plus1 length_list_update less_add_same_cancel1 less_or_eq_imp_le less_trans_Suc zero_less_numeral) + by (smt (verit) Suc_eq_plus1 length_list_update less_add_same_cancel1 less_or_eq_imp_le less_trans_Suc zero_less_numeral) qed show "length clause * (71 + 32 * (nllength \)\<^sup>2 + 2) + 1 \ ttt" using assms(1) by simp qed lemma tpsL_length: "tpsL (length clause) = tps0 [j + 1 := nltape' clause (length clause), j + 2 := nltape (map literal_n (map (rename \) (n_clause clause)))]" using tpsL_def by (simp add: n_clause_def) definition "tps9 \ tps0 [j + 1 := nltape' clause (length clause), j + 2 := (\map literal_n (map (rename \) (n_clause clause))\\<^sub>N\<^sub>L, 1)]" lemma tm9: assumes "ttt = 4 + length clause * (73 + 32 * (nllength \)\<^sup>2) + nllength (map (literal_n \ rename \) (n_clause clause))" shows "transforms tm9 (tpsL 0) ttt tps9" unfolding tm9_def proof (tform tps: tps0 tps9_def tpsL_def jk tpsL_length) show "clean_tape (tpsL (length clause) ! (j + 2))" using tpsL_def jk clean_tape_nlcontents tps0(3) by simp show "ttt = length clause * (73 + 32 * (nllength \)\<^sup>2) + 1 + (tpsL (length clause) :#: (j + 2) + 2)" using n_clause_def assms jk tpsL_length by fastforce qed lemma tm9' [transforms_intros]: assumes "ttt = 4 + 2 * length clause * (73 + 32 * (nllength \)\<^sup>2)" shows "transforms tm9 tps0 ttt tps9" proof - let ?ttt = "4 + length clause * (73 + 32 * (nllength \)\<^sup>2) + nllength (map (literal_n \ rename \) (n_clause clause))" have "?ttt \ 4 + length clause * (73 + 32 * (nllength \)\<^sup>2) + length clause * Suc (nllength \)" using clause nllength_rename by simp also have "... \ 4 + length clause * (73 + 32 * (nllength \)\<^sup>2) + length clause * (Suc (nllength \ ^ 2))" by (simp add: linear_le_pow) also have "... \ 4 + length clause * (73 + 32 * (nllength \)\<^sup>2) + length clause * (73 + nllength \ ^ 2)" by (metis One_nat_def Suc_eq_plus1 Suc_leI add.commute add_left_mono mult_le_mono2 zero_less_numeral) also have "... \ 4 + length clause * (73 + 32 * (nllength \)\<^sup>2) + length clause * (73 + 32 * nllength \ ^ 2)" by simp also have "... = 4 + 2 * length clause * (73 + 32 * (nllength \)\<^sup>2)" by simp finally have "?ttt \ 4 + 2 * length clause * (73 + 32 * (nllength \)\<^sup>2)" . then show ?thesis using tm9 assms transforms_monotone tpsL_eq_tps0 by fastforce qed definition "tps10 \ tps0 [j + 1 := (\[]\\<^sub>N\<^sub>L, 1), j + 2 := (\map literal_n (map (rename \) (n_clause clause))\\<^sub>N\<^sub>L, 1)]" lemma tm10: assumes "ttt = 11 + 2 * length clause * (73 + 32 * (nllength \)\<^sup>2) + 3 * nllength clause" shows "transforms tm10 tps0 ttt tps10" unfolding tm10_def proof (tform tps: tps0 tps9_def jk) show "tps9 ::: (j + 1) = \numlist clause\" using tps9_def jk tps0(2) nlcontents_def by simp show "proper_symbols (numlist clause)" using proper_symbols_numlist by simp show "tps10 = tps9[j + 1 := (\[]\, 1)]" by (simp add: jk nlcontents_def tps0 tps10_def tps9_def list_update_swap numlist_Nil) show "ttt = 4 + 2 * length clause * (73 + 32 * (nllength \)\<^sup>2) + (tps9 :#: (j + 1) + 2 * length (numlist clause) + 6)" using assms tps9_def jk nllength_def by simp qed lemma tm10': assumes "ttt = 11 + 64 * nllength clause * (3 + (nllength \)\<^sup>2)" shows "transforms tm10 tps0 ttt tps10" proof - let ?ttt = "11 + 2 * length clause * (73 + 32 * (nllength \)\<^sup>2) + 3 * nllength clause" have "?ttt \ 11 + 2 * nllength clause * (73 + 32 * (nllength \)\<^sup>2) + 3 * nllength clause" by (simp add: length_le_nllength) also have "... \ 11 + 2 * nllength clause * (73 + 32 * (nllength \)\<^sup>2) + 2 * 2 * nllength clause" by simp also have "... = 11 + 2 * nllength clause * (75 + 32 * (nllength \)\<^sup>2)" by algebra also have "... \ 11 + 2 * nllength clause * (96 + 32 * (nllength \)\<^sup>2)" by simp also have "... = 11 + 2 * 32 * nllength clause * (3 + (nllength \)\<^sup>2)" by simp also have "... = 11 + 64 * nllength clause * (3 + (nllength \)\<^sup>2)" by simp finally have "?ttt \ 11 + 64 * nllength clause * (3 + (nllength \)\<^sup>2)" . then show ?thesis using tm10 assms transforms_monotone by blast qed end end (* locale turing_machine_relabel-clause *) lemma transforms_tm_relabel_clauseI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k :: nat and \ clause :: "nat list" assumes "0 < j" "j + 4 < k" "length tps = k" and "\v\set clause. v div 2 < length \" assumes "tps ! j = (\\\\<^sub>N\<^sub>L, 1)" "tps ! (j + 1) = (\clause\\<^sub>N\<^sub>L, 1)" "tps ! (j + 2) = (\[]\\<^sub>N\<^sub>L, 1)" "tps ! (j + 3) = (\0\\<^sub>N, 1)" "tps ! (j + 4) = (\0\\<^sub>N, 1)" assumes "ttt = 11 + 64 * nllength clause * (3 + (nllength \)\<^sup>2)" assumes "tps' = tps [j + 1 := (\[]\\<^sub>N\<^sub>L, 1), j + 2 := (\clause_n (map (rename \) (n_clause clause))\\<^sub>N\<^sub>L, 1)]" shows "transforms (tm_relabel_clause j) tps ttt tps'" proof - interpret loc: turing_machine_relabel_clause j . show ?thesis using assms loc.tm10_eq_tm_relabel_clause loc.tps10_def loc.tm10' clause_n_def by simp qed subsection \Relabeling CNF formulas\ text \ A Turing machine can relabel a CNF formula by relabeling each clause using the TM @{const tm_relabel_clause}. The next TM accepts a CNF formula as a list of lists of literals on tape $j_1$ and a substitution $\sigma$ as a list of numbers on tape $j_2 + 1$. It outputs the relabeled formula on tape $j_2$, which initially must be empty. \ definition tm_relabel :: "tapeidx \ tapeidx \ machine" where "tm_relabel j1 j2 \ WHILE [] ; \rs. rs ! j1 \ \ DO tm_nextract \ j1 (j2 + 2) ;; tm_relabel_clause (j2 + 1) ;; tm_appendl j2 (j2 + 3) ;; tm_erase_cr (j2 + 3) DONE ;; tm_cr j1 ;; tm_cr j2" lemma tm_relabel_tm: assumes "G \ 6" and "j2 + 5 < k" and "0 < j1" and "j1 < j2" shows "turing_machine k G (tm_relabel j1 j2)" unfolding tm_relabel_def using assms tm_cr_tm tm_nextract_tm tm_appendl_tm tm_relabel_clause_tm Nil_tm tm_erase_cr_tm turing_machine_loop_turing_machine by simp locale turing_machine_relabel = fixes j1 j2 :: tapeidx begin definition "tmL1 \ tm_nextract \ j1 (j2 + 2)" definition "tmL2 \ tmL1 ;; tm_relabel_clause (j2 + 1)" definition "tmL3 \ tmL2 ;; tm_appendl j2 (j2 + 3)" definition "tmL4 \ tmL3 ;; tm_erase_cr (j2 + 3)" definition "tmL \ WHILE [] ; \rs. rs ! j1 \ \ DO tmL4 DONE" definition "tm2 \ tmL ;; tm_cr j1" definition "tm3 \ tm2 ;; tm_cr j2" lemma tm3_eq_tm_relabel: "tm3 = tm_relabel j1 j2" unfolding tm_relabel_def tm2_def tmL_def tmL4_def tmL3_def tmL2_def tmL1_def tm3_def by simp context fixes tps0 :: "tape list" and k :: nat and \ :: "nat list" and \ :: formula assumes variables: "variables \ \ {..}" assumes jk: "0 < j1" "j1 < j2" "j2 + 5 < k" "length tps0 = k" assumes tps0: "tps0 ! j1 = (\formula_n \\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! j2 = (\[]\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j2 + 1) = (\\\\<^sub>N\<^sub>L, 1)" "tps0 ! (j2 + 2) = (\[]\\<^sub>N\<^sub>L, 1)" "tps0 ! (j2 + 3) = (\[]\\<^sub>N\<^sub>L, 1)" "tps0 ! (j2 + 4) = (\0\\<^sub>N, 1)" "tps0 ! (j2 + 5) = (\0\\<^sub>N, 1)" begin abbreviation n\ :: "nat list list" where "n\ \ formula_n \" definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j1 := nlltape' n\ t, j2 := nlltape (formula_n (take t (relabel \ \)))]" lemma tpsL_eq_tps0: "tpsL 0 = tps0" using tps0 tpsL_def formula_n_def nlllength_def numlist_Nil numlist_def numlistlist_def by (metis One_nat_def list.map(1) list.size(3) list_update_id take0) definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0 [j1 := nlltape' n\ (Suc t), j2 := nlltape (formula_n (take t (relabel \ \))), j2 + 2 := (\n\ ! t\\<^sub>N\<^sub>L, 1)]" lemma tmL1 [transforms_intros]: assumes "ttt = 12 + 2 * nllength (n\ ! t)" and "t < length \" shows "transforms tmL1 (tpsL t) ttt (tpsL1 t)" unfolding tmL1_def proof (tform tps: tps0 tpsL_def tpsL1_def jk) show "t < length n\" using assms(2) formula_n_def by simp show "tpsL1 t = (tpsL t) [j1 := nlltape' n\ (Suc t), j2 + 2 := (\n\ ! t\\<^sub>N\<^sub>L, 1)]" using tpsL1_def tpsL_def by (simp add: jk list_update_swap_less) show "ttt = 12 + 2 * nllength [] + 2 * nllength (n\ ! t)" using assms(1) by simp qed definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [j1 := nlltape' n\ (Suc t), j2 := nlltape (formula_n (take t (relabel \ \))), j2 + 2 := (\[]\\<^sub>N\<^sub>L, 1), j2 + 3 := (\clause_n (map (rename \) (n_clause (n\ ! t)))\\<^sub>N\<^sub>L, 1)]" lemma tmL2 [transforms_intros]: assumes "ttt = 23 + 2 * nllength (n\ ! t) + 64 * nllength (n\ ! t) * (3 + (nllength \)\<^sup>2)" and "t < length \" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" unfolding tmL2_def proof (tform tps: assms(2) tps0 tpsL1_def jk) show "\v\set (n\ ! t). v div 2 < length \" using variables_\ variables assms(2) by (metis formula_n_def nth_map nth_mem) have "tpsL1 t ! (j2 + (1 + 2)) = (\[]\\<^sub>N\<^sub>L, 1)" using tpsL1_def jk tps0 by (simp add: numeral_3_eq_3) then show "tpsL1 t ! (j2 + 1 + 2) = (\[]\\<^sub>N\<^sub>L, 1)" by (metis add.assoc) have "tpsL1 t ! (j2 + (1 + 3)) = (\0\\<^sub>N, 1)" using tpsL1_def jk tps0 by simp then show "tpsL1 t ! (j2 + 1 + 3) = (\0\\<^sub>N, 1)" by (metis add.assoc) have "tpsL1 t ! (j2 + (1 + 4)) = (\0\\<^sub>N, 1)" using tpsL1_def jk tps0 by simp then show "tpsL1 t ! (j2 + 1 + 4) = (\0\\<^sub>N, 1)" by (metis add.assoc) have "tpsL2 t = (tpsL1 t) [j2 + (1 + 1) := (\[]\\<^sub>N\<^sub>L, 1), j2 + (1 + 2) := (\clause_n (map (rename \) (n_clause (n\ ! t)))\\<^sub>N\<^sub>L, 1)]" using jk tps0 tpsL1_def tpsL2_def by (simp add: numeral_3_eq_3) then show "tpsL2 t = (tpsL1 t) [j2 + 1 + 1 := (\[]\\<^sub>N\<^sub>L, 1), j2 + 1 + 2 := (\clause_n (map (rename \) (n_clause (n\ ! t)))\\<^sub>N\<^sub>L, 1)]" by (metis add.assoc) show "ttt = 12 + 2 * nllength (n\ ! t) + (11 + 64 * nllength (n\ ! t) * (3 + (nllength \)\<^sup>2))" using assms(1) by simp qed definition tpsL3 :: "nat \ tape list" where "tpsL3 t \ tps0 [j1 := nlltape' n\ (Suc t), j2 := nlltape (formula_n (take t (relabel \ \)) @ [clause_n (map (rename \) (n_clause (n\ ! t)))]), j2 + 2 := (\[]\\<^sub>N\<^sub>L, 1), j2 + 3 := (\clause_n (map (rename \) (n_clause (n\ ! t)))\\<^sub>N\<^sub>L, 1)]" lemma tmL3 [transforms_intros]: assumes "ttt = 29 + 2 * nllength (n\ ! t) + 64 * nllength (n\ ! t) * (3 + (nllength \)\<^sup>2) + 2 * nllength (clause_n (map (rename \) (n_clause (n\ ! t))))" and "t < length \" shows "transforms tmL3 (tpsL t) ttt (tpsL3 t)" unfolding tmL3_def proof (tform tps: assms(2) tps0 tpsL2_def jk) show "tpsL3 t = (tpsL2 t) [j2 := nlltape (formula_n (take t (relabel \ \)) @ [clause_n (map (rename \) (n_clause (n\ ! t)))])]" unfolding tpsL3_def tpsL2_def by (simp add: list_update_swap_less[of j2]) show "ttt = 23 + 2 * nllength (n\ ! t) + 64 * nllength (n\ ! t) * (3 + (nllength \)\<^sup>2) + (7 + nlllength (formula_n (take t (relabel \ \))) - Suc (nlllength (formula_n (take t (relabel \ \)))) + 2 * nllength (clause_n (map (rename \) (n_clause (n\ ! t)))))" using assms(1) by simp qed definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [j1 := nlltape' n\ (Suc t), j2 := nlltape (formula_n (take t (relabel \ \)) @ [clause_n (map (rename \) (n_clause (n\ ! t)))]), j2 + 2 := (\[]\\<^sub>N\<^sub>L, 1)]" lemma tmL4: assumes "ttt = 36 + 2 * nllength (n\ ! t) + 64 * nllength (n\ ! t) * (3 + (nllength \)\<^sup>2) + 4 * nllength (clause_n (map (rename \) (n_clause (n\ ! t))))" and "t < length \" shows "transforms tmL4 (tpsL t) ttt (tpsL4 t)" unfolding tmL4_def proof (tform tps: assms(2) tps0 tpsL3_def jk) let ?zs = "numlist (clause_n (map (rename \) (n_clause (n\ ! t))))" show "tpsL3 t ::: (j2 + 3) = \?zs\" using tpsL3_def nlcontents_def jk by simp show "proper_symbols ?zs" using proper_symbols_numlist by simp have *: "j1 \ j2 + 3" using jk by simp have "\[]\ = \[]\\<^sub>N\<^sub>L" using nlcontents_def numlist_Nil by simp then show "tpsL4 t = (tpsL3 t)[j2 + 3 := (\[]\, 1)]" using tpsL3_def tpsL4_def tps0 jk list_update_id[of tps0 "j2+3"] by (simp add: list_update_swap[OF *] list_update_swap[of _ "j2 + 3"]) show "ttt = 29 + 2 * nllength (n\ ! t) + 64 * nllength (n\ ! t) * (3 + (nllength \)\<^sup>2) + 2 * nllength (clause_n (map (rename \) (n_clause (n\ ! t)))) + (tpsL3 t :#: (j2 + 3) + 2 * length (numlist (clause_n (map (rename \) (n_clause (n\ ! t))))) + 6)" using assms(1) tpsL3_def jk nllength_def by simp qed lemma nllength_1: assumes "t < length \" shows "nllength (n\ ! t) \ nlllength n\" using assms formula_n_def nlllength_take by (metis le_less_linear length_map less_trans not_add_less2) lemma nllength_2: assumes "t < length \" shows "nllength (clause_n (map (rename \) (n_clause (n\ ! t)))) \ nlllength (formula_n (relabel \ \))" (is "?l \ ?r") proof - have "?l = nllength (clause_n (map (rename \) (\ ! t)))" using assms by (simp add: formula_n_def n_clause_n) moreover have "clause_n (map (rename \) (\ ! t)) \ set (formula_n (relabel \ \))" using assms formula_n_def relabel_def by simp ultimately show ?thesis using member_le_nlllength_1 by fastforce qed definition "tpsL4' t \ tps0 [j1 := nlltape' n\ (Suc t), j2 := nlltape (formula_n (take (Suc t) (relabel \ \)))]" lemma tpsL4: assumes "t < length \" shows "tpsL4 t = tpsL4' t" proof - have "formula_n (take t (relabel \ \)) @ [clause_n (map (rename \) (n_clause (n\ ! t)))] = formula_n (take (Suc t) (relabel \ \))" using assms formula_n_def relabel_def by (simp add: n_clause_n take_Suc_conv_app_nth) then show ?thesis using tpsL4_def tpsL4'_def jk tps0 by (smt (verit, del_insts) Suc_1 add_Suc_right add_cancel_left_right less_SucI list_update_id list_update_swap not_less_eq numeral_1_eq_Suc_0 numeral_One) qed lemma tpsL4'_eq_tpsL: "tpsL4' t = tpsL (Suc t)" using tpsL_def tpsL4'_def by simp lemma tmL4' [transforms_intros]: assumes "ttt = 36 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nlllength (formula_n (relabel \ \))" and "t < length \" shows "transforms tmL4 (tpsL t) ttt (tpsL (Suc t))" proof - let ?ttt = "36 + 2 * nllength (n\ ! t) + 64 * nllength (n\ ! t) * (3 + (nllength \)\<^sup>2) + 4 * nllength (clause_n (map (rename \) (n_clause (n\ ! t))))" have "?ttt \ 36 + 2 * nlllength n\ + 64 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nllength (clause_n (map (rename \) (n_clause (n\ ! t))))" using nllength_1 assms(2) add_le_mono add_le_mono1 mult_le_mono1 mult_le_mono2 nat_add_left_cancel_le by metis also have "... \ 36 + 2 * nlllength n\ + 64 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nlllength (formula_n (relabel \ \))" using nllength_2 assms(2) by simp also have "... \ 36 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nlllength (formula_n (relabel \ \))" by simp finally have "?ttt \ 36 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nlllength (formula_n (relabel \ \))" . then have "transforms tmL4 (tpsL t) ttt (tpsL4 t)" using assms tmL4 transforms_monotone by blast then show ?thesis using assms(2) tpsL4'_eq_tpsL tpsL4 by simp qed lemma tmL: assumes "ttt = length \ * (38 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nlllength (formula_n (relabel \ \))) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length \))" unfolding tmL_def proof (tform) let ?t = "36 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nlllength (formula_n (relabel \ \))" show "\ read (tpsL (length \)) ! j1 \ \" proof - have "tpsL (length \) ! j1 = nlltape' n\ (length n\)" using tpsL_def jk formula_n_def by simp then show ?thesis using nlltape'_tape_read[of n\ "length n\"] tapes_at_read'[of j1 "tpsL (length \)"] tpsL_def jk by simp qed show "read (tpsL t) ! j1 \ \" if "t < length \" for t proof - have "tpsL t ! j1 = nlltape' n\ t" using tpsL_def jk by simp then show ?thesis using that formula_n_def nlltape'_tape_read[of n\ t] tapes_at_read'[of j1 "tpsL t"] tpsL_def jk by simp qed show "length \ * (?t + 2) + 1 \ ttt" using assms(1) by simp qed lemma tmL' [transforms_intros]: assumes "ttt = 107 * nlllength n\ ^ 2 * (3 + nllength \ ^ 2) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length \))" proof - let ?ttt = "length \ * (38 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * nlllength (formula_n (relabel \ \))) + 1" have "?ttt \ length \ * (38 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * (Suc (nllength \) * nlllength n\)) + 1" using nlllength_relabel_variables variables by fastforce also have "... \ length \ * (38 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * ((3 + nllength \) * nlllength n\)) + 1" proof - have "Suc (nllength \) \ 3 + nllength \" by simp then show ?thesis using add_le_mono le_refl mult_le_mono by presburger qed also have "... \ length \ * (38 + 65 * nlllength n\ * (3 + (nllength \)\<^sup>2) + 4 * ((3 + nllength \ ^ 2) * nlllength n\)) + 1" using linear_le_pow by simp also have "... = length \ * (38 + 69 * nlllength n\ * (3 + (nllength \)\<^sup>2)) + 1" by simp also have "... \ nlllength n\ * (38 + 69 * nlllength n\ * (3 + (nllength \)\<^sup>2)) + 1" using length_le_nlllength formula_n_def by (metis add_mono_thms_linordered_semiring(3) length_map mult_le_mono1) also have "... = 38 * nlllength n\ + (69 * nlllength n\ ^ 2 * (3 + (nllength \)\<^sup>2)) + 1" by algebra also have "... \ 38 * nlllength n\ ^ 2 * (3 + nllength \ ^ 2) + (69 * nlllength n\ ^ 2 * (3 + (nllength \)\<^sup>2)) + 1" proof - have "nlllength n\ \ nlllength n\ ^ 2 * (3 + nllength \ ^ 2)" using linear_le_pow by (metis One_nat_def Suc_leI add_gr_0 mult_le_mono nat_mult_1_right zero_less_numeral) then show ?thesis by simp qed also have "... = 107 * nlllength n\ ^ 2 * (3 + nllength \ ^ 2) + 1" by simp finally have "?ttt \ 107 * nlllength n\ ^ 2 * (3 + nllength \ ^ 2) + 1" . then show ?thesis using tmL assms transforms_monotone by blast qed definition tps1 :: "tape list" where "tps1 \ tps0 [j1 := nlltape' n\ (length \), j2 := nlltape (formula_n (relabel \ \))]" lemma tps1_eq_tpsL: "tps1 = tpsL (length \)" using tps1_def tpsL_def jk tps0 relabel_def by simp definition "tps2 \ tps0 [j2 := nlltape (formula_n (relabel \ \))]" lemma tm2 [transforms_intros]: assumes "ttt = Suc (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + Suc (Suc (Suc (nlllength n\)))" shows "transforms tm2 (tpsL 0) ttt tps2" unfolding tm2_def proof (tform tps: tps0 tpsL_def tps1_def jk) have *: "tpsL (length \) ! j1 = nlltape' n\ (length \)" using tpsL_def jk by simp then show "clean_tape (tpsL (length \) ! j1)" using clean_tape_nllcontents by simp have "tpsL (length \) ! j1 |#=| 1 = nlltape' n\ 0" using * by simp then show "tps2 = (tpsL (length \))[j1 := tpsL (length \) ! j1 |#=| 1]" using tps0 jk tps2_def tps1_eq_tpsL tps1_def by (metis (no_types, lifting) One_nat_def list_update_id list_update_overwrite list_update_swap_less nlllength_Nil take0) show "ttt = 107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2) + 1 + (tpsL (length \) :#: j1 + 2)" using assms tpsL_def jk formula_n_def by simp qed definition "tps3 \ tps0 [j2 := nlltape' (formula_n (relabel \ \)) 0]" lemma tm3: assumes "ttt = 7 + (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + nlllength n\ + nlllength (formula_n (relabel \ \))" shows "transforms tm3 (tpsL 0) ttt tps3" unfolding tm3_def proof (tform tps: assms tps0 tps2_def tps3_def jk) show "clean_tape (tps2 ! j2)" using tps2_def jk clean_tape_nllcontents by simp qed lemma tm3' [transforms_intros]: assumes "ttt = 7 + (108 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2))" shows "transforms tm3 tps0 ttt tps3" proof - let ?ttt = "7 + (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + nlllength n\ + nlllength (formula_n (relabel \ \))" have "?ttt \ 7 + (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + nlllength n\ + Suc (nllength \) * nlllength n\" using variables nlllength_relabel_variables by simp also have "... = 7 + (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + (2 + nllength \) * nlllength n\" by simp also have "... \ 7 + (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + (2 + nllength \ ^ 2) * nlllength n\" using linear_le_pow by simp also have "... \ 7 + (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + (3 + nllength \ ^ 2) * nlllength n\" by (metis add_2_eq_Suc add_Suc_shift add_mono_thms_linordered_semiring(2) le_add2 mult.commute mult_le_mono2 numeral_3_eq_3) also have "... \ 7 + (107 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2)) + (3 + nllength \ ^ 2) * nlllength n\ ^ 2" using linear_le_pow by simp also have "... = 7 + (108 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2))" by simp finally have "?ttt \ 7 + (108 * (nlllength n\)\<^sup>2 * (3 + (nllength \)\<^sup>2))" . then show ?thesis using tm3 assms tpsL_eq_tps0 transforms_monotone by simp qed end (* context tps0 *) end (* locale turing_machine_relabel *) lemma transforms_tm_relabelI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and ttt k :: nat and \ :: "nat list" and \ :: formula assumes "0 < j1" and "j1 < j2" and "j2 + 5 < k" and "length tps = k" and "variables \ \ {..}" assumes "tps ! j1 = (\formula_n \\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps ! j2 = (\[]\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps ! (j2 + 1) = (\\\\<^sub>N\<^sub>L, 1)" "tps ! (j2 + 2) = (\[]\\<^sub>N\<^sub>L, 1)" "tps ! (j2 + 3) = (\[]\\<^sub>N\<^sub>L, 1)" "tps ! (j2 + 4) = (\0\\<^sub>N, 1)" "tps ! (j2 + 5) = (\0\\<^sub>N, 1)" assumes "tps' = tps [j2 := nlltape' (formula_n (relabel \ \)) 0]" assumes "ttt = 7 + (108 * (nlllength (formula_n \))\<^sup>2 * (3 + (nllength \)\<^sup>2))" shows "transforms (tm_relabel j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_relabel j1 j2 . show ?thesis using assms loc.tm3_eq_tm_relabel loc.tm3' loc.tps3_def by simp qed section \Listing the head positions of a Turing machine\ text \ The formulas $\chi_t$ used for $\Phi_9$ require the functions $\inputpos$ and $\prev$, or more precisely the substitutions $\rho_t$ and $\rho'_t$ do. In this section we build a TM that simulates a two-tape TM $M$ on some input until it halts. During the simulation it creates two lists: one with the sequence of head positions on $M$'s input tape and one with the sequence of head positions on $M$'s output tape. The first list directly provides $\inputpos$; the second list allows the computation of $\prev$ using the TM @{const tm_prev}. \ subsection \Simulating and logging head movements\ text \ At the core of the simulation is the following Turing command. It interprets the tapes $j + 7$ and $j + 8$ as input tape and output tape of a two-tape Turing machine $M$ and the symbol in the first cell of tape $j + 6$ as the state of $M$. It then applies the actions of $M$ in this configuration to the tapes $j + 7$ and $j + 8$ and adapts the state on tape $j + 6$ accordingly. On top of that it writes (``logs'') to tape $j$ the direction in which $M$'s input tape head has moved and to tape $j + 3$ the direction in which $M$'s work tape head has moved. The head movement directions are encoded by the symbols $\Box$, $\triangleright$, and \textbf{0} for Left, Stay, and Right, respectively. The command is parameterized by the TM $M$ and its alphabet size $G$ (and as usual the tape index $j$). The command does nothing if the state on tape $j + 6$ is the halting state or if the symbol read from the simulated tape $j + 7$ or $j + 8$ is outside the alphabet $G$. \null \ definition direction_to_symbol :: "direction \ symbol" where "direction_to_symbol d \ (case d of Left \ \ | Stay \ \ | Right \ \)" lemma direction_to_symbol_less: "direction_to_symbol d < 3" using direction_to_symbol_def by (cases d) simp_all definition cmd_simulog :: "nat \ machine \ tapeidx \ command" where "cmd_simulog G M j rs \ (1, if rs ! (j + 6) \ length M \ rs ! (j + 7) \ G \ rs ! (j + 8) \ G then map (\z. (z, Stay)) rs else map (\i. let sas = (M ! (rs ! (j + 6))) [rs ! (j + 7), rs ! (j + 8)] in if i = j then (direction_to_symbol (sas [~] 0), Stay) else if i = j + 3 then (direction_to_symbol (sas [~] 1), Stay) else if i = j + 6 then (fst sas, Stay) else if i = j + 7 then sas [!] 0 else if i = j + 8 then sas [!] 1 else (rs ! i, Stay)) [0.. j + 9" and "j > 0" and "H \ Suc (length M)" and "H \ G" shows "turing_command k 1 H (cmd_simulog G M j)" proof show "\gs. length gs = k \ length ([!!] cmd_simulog G M j gs) = length gs" using cmd_simulog_def by simp have G: "H \ 4" using assms(1,5) turing_machine_def by simp show "cmd_simulog G M j rs [.] i < H" if "length rs = k" and "(\i. i < length rs \ rs ! i < H)" and "i < length rs" for rs i proof (cases "rs ! (j + 6) \ length M \ rs ! (j + 7) \ G \ rs ! (j + 8) \ G") case True then show ?thesis using assms that cmd_simulog_def by simp next case False then have inbounds: "rs ! (j + 6) < length M" by simp let ?cmd = "M ! (rs ! (j + 6))" let ?gs = "[rs ! (j + 7), rs ! (j + 8)]" let ?sas = "?cmd ?gs" have lensas: "length (snd ?sas) = 2" using assms(1) inbounds turing_commandD(1) by (metis length_Cons list.size(3) numeral_2_eq_2 turing_machineD(3)) consider "i = j" | "i = j + 3" | "i = j + 6" | "i = j + 7" | "i = j + 8" | "i \ j \ i \ j + 3 \ i \ j + 6 \ i \ j + 7 \ i \ j + 8" by linarith then show ?thesis proof (cases) case 1 then have "cmd_simulog G M j rs [!] i = (direction_to_symbol (?sas [~] 0), Stay)" using cmd_simulog_def False that by simp then have "cmd_simulog G M j rs [.] i < 3" using direction_to_symbol_less by simp then show ?thesis using G by simp next case 2 then have "cmd_simulog G M j rs [!] i = (direction_to_symbol (?sas [~] 1), Stay)" using cmd_simulog_def False that by simp then have "cmd_simulog G M j rs [.] i < 3" using direction_to_symbol_less by simp then show ?thesis using G by simp next case 3 then have "cmd_simulog G M j rs [!] i = (fst ?sas, Stay)" using cmd_simulog_def False that by simp then have "cmd_simulog G M j rs [.] i \ length M" using assms inbounds turing_commandD(4) turing_machineD(3) by (metis One_nat_def Suc_1 fst_conv length_Cons list.size(3)) then show ?thesis using assms(4) by simp next case 4 then have "cmd_simulog G M j rs [!] i = ?sas [!] 0" using cmd_simulog_def False that by simp then show ?thesis using 4 assms inbounds turing_machine_def that lensas turing_commandD(3) by (metis One_nat_def Suc_1 length_Cons list.size(3) nth_Cons_0 nth_mem zero_less_numeral) next case 5 then have *: "cmd_simulog G M j rs [!] i = ?sas [!] 1" using cmd_simulog_def False that by simp have "turing_command 2 (length M) G ?cmd" using assms(1) inbounds turing_machine_def by simp moreover have "symbols_lt G ?gs" using False less_2_cases_iff numeral_2_eq_2 by simp ultimately have "?sas [.] 1 < G" using turing_commandD(2) by simp then show ?thesis using assms * by simp next case 6 then have "cmd_simulog G M j rs [!] i = (rs ! i, Stay)" using cmd_simulog_def False that(3) by simp then show ?thesis using that by simp qed qed show "cmd_simulog G M j rs [.] 0 = rs ! 0" if "length rs = k" and "0 < k" for rs proof (cases "rs ! (j + 6) \ length M \ rs ! (j + 7) \ G \ rs ! (j + 8) \ G") case True then show ?thesis using assms that cmd_simulog_def by simp next case False then show ?thesis using assms that cmd_simulog_def by simp qed show "\gs. length gs = k \ [*] (cmd_simulog G M j gs) \ 1" using cmd_simulog_def by simp qed text \ The logging Turing machine consists only of the logging command. \ definition tm_simulog :: "nat \ machine \ tapeidx \ machine" where "tm_simulog G M j \ [cmd_simulog G M j]" lemma tm_simulog_tm: fixes H :: nat assumes "turing_machine 2 G M" and "k \ j + 9" and "j > 0" and "H \ Suc (length M)" and "H \ G" shows "turing_machine k H (tm_simulog G M j)" using turing_command_cmd_simulog tm_simulog_def assms turing_machine_def by simp lemma transforms_tm_simulogI [transforms_intros]: fixes G :: nat and M :: machine and j :: tapeidx fixes k :: nat and tps tps' :: "tape list" and xs :: "symbol list" assumes "turing_machine 2 G M" and "k \ j + 9" and "j > 0" and "symbols_lt G xs" and "cfg = execute M (start_config 2 xs) t" and "fst cfg < length M" and "length tps = k" assumes "tps ! j = \dummy1\" "tps ! (j + 3) = \dummy2\" "tps ! (j + 6) = \fst cfg\" "tps ! (j + 7) = cfg 0" "tps ! (j + 8) = cfg 1" assumes "tps' = tps [j := \direction_to_symbol ((M ! (fst cfg)) (config_read cfg) [~] 0)\, j + 3 := \direction_to_symbol ((M ! (fst cfg)) (config_read cfg) [~] 1)\, j + 6 := \fst (execute M (start_config 2 xs) (Suc t))\, j + 7 := execute M (start_config 2 xs) (Suc t) 0, j + 8 := execute M (start_config 2 xs) (Suc t) 1]" shows "transforms (tm_simulog G M j) tps 1 tps'" proof - have "sem (cmd_simulog G M j) (0, tps) = (1, tps')" proof (rule semI) define H where "H = max G (Suc (length M))" then have "H \ Suc (length M)" "H \ G" by simp_all then show "proper_command k (cmd_simulog G M j)" using assms cmd_simulog_def by simp show "length tps = k" and "length tps' = k" using assms by simp_all show "fst (cmd_simulog G M j (read tps)) = 1" using cmd_simulog_def sem' by simp define rs where "rs = read tps" then have lenrs: "length rs = k" by (simp add: assms rs_def read_length) have rs6: "rs ! (j + 6) = fst cfg" using rs_def tapes_at_read'[of "j + 6" tps] assms by simp then have inbounds: "rs ! (j + 6) < length M" using assms by simp have rs7: "rs ! (j + 7) = cfg <.> 0" using rs_def tapes_at_read'[of "j + 7" tps] assms by simp then have rs7_less: "rs ! (j + 7) < G" using assms tape_alphabet[OF assms(1,4)] by simp have rs8: "rs ! (j + 8) = cfg <.> 1" using rs_def tapes_at_read'[of "j + 8" tps] assms by simp then have rs8_less: "rs ! (j + 8) < G" using assms tape_alphabet[OF assms(1,4)] by simp let ?gs = "[rs ! (j + 7), rs ! (j + 8)]" have gs: "?gs = config_read cfg" proof (rule nth_equalityI) show "length [rs ! (j + 7), rs ! (j + 8)] = length (config_read cfg)" using assms execute_num_tapes start_config_length read_length by simp then show "[rs ! (j + 7), rs ! (j + 8)] ! i = config_read cfg ! i" if "i < length [rs ! (j + 7), rs ! (j + 8)]" for i using assms that rs7 rs8 read_length tapes_at_read' by (metis One_nat_def length_Cons less_2_cases_iff list.size(3) nth_Cons_0 nth_Cons_Suc numeral_2_eq_2) qed let ?cmd = "M ! (rs ! (j + 6))" let ?sas = "?cmd ?gs" have lensas: "length (snd ?sas) = 2" using assms(1) inbounds turing_commandD(1) turing_machine_def by (metis One_nat_def Suc_1 canrepr_1 list.size(4) nlength_1_simp nth_mem plus_1_eq_Suc) have sas: "?sas = (M ! (fst cfg)) (config_read cfg)" using rs6 gs by simp have "act (cmd_simulog G M j rs [!] i) (tps ! i) = tps' ! i" if "i < k" for i proof - have "cmd_simulog G M j rs = (1, map (\i. let sas = (M ! (rs ! (j + 6))) [rs ! (j + 7), rs ! (j + 8)] in if i = j then (direction_to_symbol (sas [~] 0), Stay) else if i = j + 3 then (direction_to_symbol (sas [~] 1), Stay) else if i = j + 6 then (fst sas, Stay) else if i = j + 7 then sas [!] 0 else if i = j + 8 then sas [!] 1 else (rs ! i, Stay)) [0.. j \ i \ j + 3 \ i \ j + 6 \ i \ j + 7 \ i \ j + 8" by linarith then show ?thesis proof (cases) case 1 then have "cmd_simulog G M j rs [!] i = (direction_to_symbol (?sas [~] 0), Stay)" using * by simp moreover have "tps' ! i = \direction_to_symbol (?sas [~] 0)\" using 1 assms sas by simp ultimately show ?thesis using 1 act_onesie assms(8) by simp next case 2 then have "cmd_simulog G M j rs [!] i = (direction_to_symbol (?sas [~] 1), Stay)" using * by simp moreover have "tps' ! i = \direction_to_symbol (?sas [~] 1)\" using 2 assms sas by simp ultimately show ?thesis using 2 act_onesie assms by simp next case 3 then have "cmd_simulog G M j rs [!] i = (fst ?sas, Stay)" using * by simp moreover have "tps' ! i = \fst (execute M (start_config 2 xs) (Suc t))\" using 3 assms by simp ultimately show ?thesis using 3 act_onesie assms by (metis exe_lt_length execute.simps(2) sas sem_fst) next case 4 then have "cmd_simulog G M j rs [!] i = ?sas [!] 0" using * by simp moreover have "tps' ! i = execute M (start_config 2 xs) (Suc t) 0" using 4 assms by simp moreover have "proper_command 2 (M ! (rs ! (j + 6)))" using assms(1,6) rs6 turing_machine_def turing_commandD(1) turing_machineD by metis ultimately show ?thesis using 4 assms(1,11,5,6) exe_lt_length gs read_length rs6 sem_snd turing_machine_def by (metis execute.simps(2) length_Cons list.size(3) numeral_2_eq_2 zero_less_numeral) next case 5 then have "cmd_simulog G M j rs [!] i = ?sas [!] 1" using * by simp moreover have "tps' ! i = execute M (start_config 2 xs) (Suc t) 1" using 5 assms by simp moreover have "proper_command 2 (M ! (rs ! (j + 6)))" using assms(1,6) rs6 turing_machineD turing_commandD(1) by metis ultimately show ?thesis using 5 assms(1,12,5,6) exe_lt_length gs read_length rs6 sem_snd turing_machine_def by (metis One_nat_def execute.simps(2) length_Cons less_2_cases_iff list.size(3) numeral_2_eq_2) next case 6 then have "cmd_simulog G M j rs [!] i = (rs ! i, Stay)" using * by simp moreover have "tps' ! i = tps ! i" using 6 assms(13) by simp ultimately show ?thesis using 6 assms act_Stay rs_def that by metis qed qed then show "act (cmd_simulog G M j (read tps) [!] i) (tps ! i) = tps' ! i" if "i < k" for i using that rs_def by simp qed moreover have "execute (tm_simulog G M j) (0, tps) 1 = sem (cmd_simulog G M j) (0, tps)" using tm_simulog_def by (simp add: exe_lt_length) ultimately have "execute (tm_simulog G M j) (0, tps) 1 = (1, tps')" by simp then show ?thesis using transforms_def transits_def tm_simulog_def by auto qed subsection \Adjusting head position counters\ text \ The Turing machine @{const tm_simulog} logs the head movements, but what we need is a list of all the head positions during the execution of $M$. The next TM maintains a number for a head position and adjusts it based on a head movement symbol as provided by @{const tm_simulog}. More precisely, the next Turing machine accepts on tape $j$ a symbol encoding a direction, on tape $j + 1$ a number representing a head position, and on tape $j + 2$ a list of numbers. Depending on the symbol on tape $j$ it decreases, increases or leaves unchanged the number on tape $j + 1$. Then it appends this adjusted number to the list on tape $j + 2$. \ definition tm_adjust_headpos :: "nat \ machine" where "tm_adjust_headpos j \ IF \rs. rs ! j = \ THEN tm_decr (j + 1) ELSE IF \rs. rs ! j = \ THEN tm_incr (j + 1) ELSE [] ENDIF ENDIF ;; tm_append (j + 2) (j + 1)" lemma tm_adjust_headpos_tm: assumes "G \ 5" and "j + 2 < k" shows "turing_machine k G (tm_adjust_headpos j)" unfolding tm_adjust_headpos_def using assms turing_machine_branch_turing_machine tm_decr_tm tm_incr_tm tm_append_tm Nil_tm turing_machine_sequential_turing_machine by simp locale turing_machine_adjust_headpos = fixes j :: tapeidx begin definition "tm1 \ IF \rs. rs ! j = \ THEN tm_incr (j + 1) ELSE [] ENDIF" definition "tm2 \ IF \rs. rs ! j = \ THEN tm_decr (j + 1) ELSE tm1 ENDIF" definition "tm3 \ tm2 ;; tm_append (j + 2) (j + 1)" lemma tm3_eq_tm_adjust_headpos: "tm3 = tm_adjust_headpos j" unfolding tm1_def tm2_def tm3_def tm_adjust_headpos_def by simp context fixes tps :: "tape list" and jj :: tapeidx and k t :: nat and xs :: "symbol list" fixes M :: machine fixes G cfg assumes jk: "length tps = k" "k \ j + 3" "jj < 2" assumes M: "turing_machine 2 G M" assumes xs: "symbols_lt G xs" assumes cfg: "cfg = execute M (start_config 2 xs) t" "fst cfg < length M" assumes tps0: "tps ! j = \direction_to_symbol ((M ! (fst cfg)) (config_read cfg) [~] jj)\" "tps ! (j + 1) = (\cfg <#> jj\\<^sub>N, 1)" "tps ! (j + 2) = nltape (map (\t. (execute M (start_config 2 xs) t <#> jj)) [0.. k" using jk by simp abbreviation exc :: "symbol list \ nat \ config" where "exc y tt \ execute M (start_config 2 y) tt" lemma read_tps_j: "read tps ! j = direction_to_symbol ((M ! (fst cfg)) (config_read cfg) [~] jj)" using tps0 onesie_read jk tapes_at_read' by (metis less_add_same_cancel1 less_le_trans zero_less_numeral) lemma write_symbol: "\v. execute M (start_config 2 xs) (Suc t) jj = act (v, (M ! (fst cfg)) (config_read cfg) [~] jj) (cfg jj)" proof - let ?d = "(M ! (fst cfg)) (config_read cfg) [~] jj" obtain v where v: "(M ! (fst cfg)) (config_read cfg) [!] jj = (v, ?d)" by (simp add: prod_eq_iff) have "execute M (start_config 2 xs) (Suc t) jj = exe M cfg jj" using cfg(1) by simp also have "... = sem (M ! (fst cfg)) cfg jj" by (simp add: cfg(2) exe_lt_length) also have "... = act ((M ! (fst cfg)) (config_read cfg) [!] jj) (cfg jj)" using sem_snd_tm M cfg execute_num_tapes start_config_length jk by (metis (no_types, lifting) numeral_2_eq_2 prod.exhaust_sel zero_less_Suc) also have "... = act (v, ?d) (cfg jj)" using v by simp finally have *: "execute M (start_config 2 xs) (Suc t) jj = act (v, ?d) (cfg jj)" . then show ?thesis by auto qed lemma tm1 [transforms_intros]: assumes "ttt = 7 + 2 * nlength (cfg <#> jj)" and "(M ! (fst cfg)) (config_read cfg) [~] jj \ Left" and "tps' = tps[j + 1 := (\execute M (start_config 2 xs) (Suc t) <#> jj\\<^sub>N, 1)]" shows "transforms tm1 tps ttt tps'" unfolding tm1_def proof (tform) let ?d = "(M ! (fst cfg)) (config_read cfg) [~] jj" obtain v where v: "execute M (start_config 2 xs) (Suc t) jj = act (v, ?d) (cfg jj)" using write_symbol by auto { assume "read tps ! j = 2" then have "?d = Right" using read_tps_j assms(2) direction_to_symbol_def by (cases ?d) simp_all show "j + 1 < length tps" using jk by simp show "tps ! (j + 1) = (\cfg <#> jj\\<^sub>N, 1)" using tps0 by simp_all show "tps' = tps[j + 1 := (\Suc (cfg <#> jj)\\<^sub>N, 1)]" proof - have "execute M (start_config 2 xs) (Suc t) jj = cfg jj |:=| v |+| 1" using v `?d = Right` act_Right' by simp then have "execute M (start_config 2 xs) (Suc t) <#> jj = cfg <#> jj + 1" by simp then show ?thesis using assms(3) by simp qed } { assume "read tps ! j \ 2" then have "?d = Stay" using read_tps_j assms(2) direction_to_symbol_def by (cases ?d) simp_all then have "execute M (start_config 2 xs) (Suc t) jj = cfg jj |:=| v" using v act_Stay' by simp then have "execute M (start_config 2 xs) (Suc t) <#> jj = cfg <#> jj" by simp then show "tps' = tps" using assms(3) tps0 by (metis list_update_id) } show "(5 + 2 * nlength (cfg <#> jj)) + 2 \ ttt" "0 + 1 \ ttt" using assms(1) by simp_all qed lemma tm2 [transforms_intros]: assumes "ttt = 10 + 2 * nlength (cfg <#> jj)" and "tps' = tps[j + 1 := (\execute M (start_config 2 xs) (Suc t) <#> jj\\<^sub>N, 1)]" shows "transforms tm2 tps ttt tps'" unfolding tm2_def proof (tform tps: k_ge_2 jk assms) let ?d = "(M ! (fst cfg)) (config_read cfg) [~] jj" show "8 + 2 * nlength (cfg <#> jj) + 2 \ ttt" using assms(1) by simp show "read tps ! j \ \ \ ?d \ Left" using read_tps_j direction_to_symbol_def by (cases ?d) simp_all { assume 0: "read tps ! j = \" show "tps ! (j + 1) = (\cfg <#> jj\\<^sub>N, 1)" using tps0 by simp show "tps' = tps[j + 1 := (\cfg <#> jj - 1\\<^sub>N, 1)]" proof - let ?d = "(M ! (fst cfg)) (config_read cfg) [~] jj" obtain v where v: "execute M (start_config 2 xs) (Suc t) jj = act (v, ?d) (cfg jj)" using write_symbol by auto then have "?d = Left" using 0 read_tps_j assms(2) direction_to_symbol_def by (cases ?d) simp_all then have "execute M (start_config 2 xs) (Suc t) jj = cfg jj |:=| v |-| 1" using v act_Left' by simp then have "execute M (start_config 2 xs) (Suc t) <#> jj = cfg <#> jj - 1" by simp then show ?thesis using assms(2) by simp qed } qed lemma tm3: assumes "ttt = 16 + 2 * nlength (cfg <#> jj) + 2 * nlength (exc xs (Suc t) <#> jj)" and "tps' = tps [j + 1 := (\execute M (start_config 2 xs) (Suc t) <#> jj\\<^sub>N, 1), j + 2 := nltape (map (\t. (execute M (start_config 2 xs) t <#> jj)) [0..?ns\\<^sub>N\<^sub>L, ?i)" using tps0 by simp show "Suc (nllength ?ns) \ ?i" by simp show "tps' = tps [j + 1 := (\?n\\<^sub>N, 1), j + 2 := nltape (?ns @ [snd (exe M (exc xs t)) :#: jj])]" using assms(2) nlcontents_def nllength_def by simp show "ttt = 10 + 2 * nlength (cfg <#> jj) + (7 + nllength (map (\t. exc xs t <#> jj) [0..t. exc xs t <#> jj) [0.. j + 3" and "jj < 2" and "symbols_lt G xs" and cfg: "cfg = execute M (start_config 2 xs) t" "fst cfg < length M" assumes "tps ! j = \direction_to_symbol ((M ! (fst cfg)) (config_read cfg) [~] jj)\" "tps ! (j + 1) = (\cfg <#> jj\\<^sub>N, 1)" "tps ! (j + 2) = nltape (map (\t. (execute M (start_config 2 xs) t <#> jj)) [0..t. execute M (start_config 2 xs) t <#> jj \ max_head_pos" assumes ttt: "ttt = 16 + 4 * nlength max_head_pos" assumes "tps' = tps [j + 1 := (\execute M (start_config 2 xs) (Suc t) <#> jj\\<^sub>N, 1), j + 2 := nltape (map (\t. (execute M (start_config 2 xs) t <#> jj)) [0.. ttt" proof - have "?ttt \ 16 + 2 * nlength (cfg <#> jj) + 2 * nlength max_head_pos" using max_head_pos nlength_mono by (meson add_le_mono le_refl mult_le_mono2) also have "... \ 16 + 2 * nlength max_head_pos + 2 * nlength max_head_pos" using max_head_pos cfg nlength_mono by simp also have "... = 16 + 4 * nlength max_head_pos" by simp finally show ?thesis using ttt by simp qed ultimately show ?thesis using transforms_monotone by simp qed subsection \Listing the head positions\ text \ The next Turing machine is essentially a loop around the TM @{const tm_simulog}, which outputs head movements, combined with two instances of the TM @{const tm_adjust_headpos}, each of which maintains a head positions lists. The loop ends when the simulated machine reaches the halting state. If the simulated machine does not halt, neither does the simulator, but we will not consider this case when we analyze the semantics. The TM receives an input on tape $j + 7$. During the simulation of $M$ this tape is a replica of the simulated machine's input tape, and tape $j + 8$ is a replica of the work/output tape. The lists of the head positions will be on tapes $j + 2$ and $j + 5$ for the input tape and work/output tape, respectively. \ definition tm_list_headpos :: "nat \ machine \ tapeidx \ machine" where "tm_list_headpos G M j \ tm_right_many {j + 1, j + 2, j + 4, j+ 5} ;; tm_write (j + 6) \ ;; tm_append (j + 2) (j + 1) ;; tm_append (j + 5) (j + 4) ;; WHILE [] ; \rs. rs ! (j + 6) < length M DO tm_simulog G M j ;; tm_adjust_headpos j ;; tm_adjust_headpos (j + 3) ;; tm_write_many {j, j + 3} \ DONE ;; tm_write (j + 6) \ ;; tm_cr (j + 2) ;; tm_cr (j + 5)" lemma tm_list_headpos_tm: fixes H :: nat assumes "turing_machine 2 G M" and "k \ j + 9" and "j > 0" and "H \ Suc (length M)" and "H \ G" assumes "H \ 5" shows "turing_machine k H (tm_list_headpos G M j)" unfolding tm_list_headpos_def using assms turing_machine_loop_turing_machine turing_machine_sequential_turing_machine Nil_tm tm_append_tm tm_simulog_tm tm_adjust_headpos_tm tm_right_many_tm tm_write_tm tm_write_many_tm tm_cr_tm by simp locale turing_machine_list_headpos = fixes G :: nat and M :: machine and j :: tapeidx begin definition "tm1 \ tm_right_many {j + 1, j + 2, j + 4, j + 5}" definition "tm2 \ tm1 ;; tm_write (j + 6) \" definition "tm3 \ tm2 ;; tm_append (j + 2) (j + 1)" definition "tm4 \ tm3 ;; tm_append (j + 5) (j + 4)" definition "tmL1 \ tm_simulog G M j" definition "tmL2 \ tmL1 ;; tm_adjust_headpos j" definition "tmL3 \ tmL2 ;; tm_adjust_headpos (j + 3)" definition "tmL4 \ tmL3 ;; tm_write_many {j, j + 3} \" definition "tmL \ WHILE [] ; \rs. rs ! (j + 6) < length M DO tmL4 DONE" definition "tm5 \ tm4 ;; tmL" definition "tm6 \ tm5 ;; tm_write (j + 6) \" definition "tm7 \ tm6 ;; tm_cr (j + 2)" definition "tm8 \ tm7 ;; tm_cr (j + 5)" lemma tm8_eq_tm_list_headpos: "tm8 = tm_list_headpos G M j" unfolding tm1_def tm2_def tm3_def tm4_def tmL1_def tmL2_def tmL3_def tmL4_def tmL_def tm5_def tm6_def tm7_def tm8_def tm_list_headpos_def by simp context fixes tps0 :: "tape list" fixes thalt k :: nat and xs :: "symbol list" assumes M: "turing_machine 2 G M" assumes jk: "k \ j + 9" "j > 0" "length tps0 = k" assumes thalt: "\t\\" "tps0 ! (j + 1) = (\0\\<^sub>N, 0)" "tps0 ! (j + 2) = (\[]\\<^sub>N\<^sub>L, 0)" "tps0 ! (j + 3) = \\\" "tps0 ! (j + 4) = (\0\\<^sub>N, 0)" "tps0 ! (j + 5) = (\[]\\<^sub>N\<^sub>L, 0)" "tps0 ! (j + 6) = \\\" "tps0 ! (j + 7) = (\xs\, 0)" "tps0 ! (j + 8) = (\[]\, 0)" begin abbreviation exec :: "nat \ config" where "exec t \ execute M (start_config 2 xs) t" lemma max_head_pos_0: "\t. exec t <#> 0 \ thalt" using thalt M head_pos_le_halting_time by simp lemma max_head_pos_1: "\t. exec t <#> 1 \ thalt" using thalt M head_pos_le_halting_time by simp definition "tps1 \ tps0 [(j + 1) := (\0\\<^sub>N, 1), (j + 2) := (\[]\\<^sub>N\<^sub>L, 1), (j + 4) := (\0\\<^sub>N, 1), (j + 5) := (\[]\\<^sub>N\<^sub>L, 1), (j + 6) := \\\]" lemma tm1 [transforms_intros]: assumes "ttt = 1" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: assms tps0 jk tps1_def) show "tps1 = map (\i. if i \ {j + 1, j + 2, j + 4, j + 5} then tps0 ! i |+| 1 else tps0 ! i) [0.. ?J") case True have "tps1 ! (j + 1) = ?rhs ! (j + 1)" using tps1_def jk tps0 by fastforce moreover have "tps1 ! (j + 2) = ?rhs ! (j + 2)" using tps1_def jk tps0 by fastforce moreover have "tps1 ! (j + 4) = ?rhs ! (j + 4)" using tps1_def jk tps0 by fastforce moreover have "tps1 ! (j + 5) = ?rhs ! (j + 5)" using tps1_def jk tps0 by fastforce ultimately show ?thesis using True by fast next case notinJ: False then have *: "?rhs ! i = tps0 ! i" using that len by simp show ?thesis proof (cases "i = j + 6") case True then show ?thesis using * that tps0(7) tps1_def by simp next case False then have "tps1 ! i = tps0 ! i" using tps1_def notinJ that by simp then show ?thesis using * by simp qed qed qed qed definition "tps2 \ tps0 [(j + 1) := (\0\\<^sub>N, 1), (j + 2) := (\[]\\<^sub>N\<^sub>L, 1), (j + 4) := (\0\\<^sub>N, 1), (j + 5) := (\[]\\<^sub>N\<^sub>L, 1), (j + 6) := \\\]" lemma tm2 [transforms_intros]: assumes "ttt = 2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: assms tps0 jk tps2_def tps1_def) show "tps2 = tps1[j + 6 := tps1 ! (j + 6) |:=| 0]" using tps2_def tps1_def jk onesie_write - by (smt (z3) list_update_beyond list_update_overwrite nth_list_update_eq verit_comp_simplify1(3)) + by (smt (verit) list_update_beyond list_update_overwrite nth_list_update_eq verit_comp_simplify1(3)) qed definition "tps3 \ tps0 [(j + 1) := (\0\\<^sub>N, 1), (j + 2) := nltape [0], (j + 4) := (\0\\<^sub>N, 1), (j + 5) := (\[]\\<^sub>N\<^sub>L, 1), (j + 6) := \\\]" lemma tm3 [transforms_intros]: assumes "ttt = 8" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps0 jk tps2_def tps3_def) show "tps3 = tps2[j + 2 := nltape ([] @ [0])]" using tps3_def jk tps2_def by (smt (verit, ccfv_SIG) Suc3_eq_add_3 add_2_eq_Suc add_less_cancel_left lessI list_update_overwrite list_update_swap not_add_less2 numeral_2_eq_2 numeral_3_eq_3 numeral_Bit0 numeral_plus_numeral self_append_conv2 semiring_norm(4) semiring_norm(5)) show "ttt = 2 + (7 + nllength [] - Suc 0 + 2 * nlength 0)" using assms by simp qed definition "tps4 \ tps0 [(j + 1) := (\0\\<^sub>N, 1), (j + 2) := nltape [0], (j + 4) := (\0\\<^sub>N, 1), (j + 5) := nltape [0], (j + 6) := \\\]" lemma tm4 [transforms_intros]: assumes "ttt = 14" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: tps3_def jk tps0 tps4_def) show "tps4 = tps3[j + 5 := nltape ([] @ [0])]" unfolding tps3_def tps4_def using jk tps0 - by (smt (z3) add_left_imp_eq list_update_overwrite list_update_swap num.simps(8) numeral_eq_iff self_append_conv2) + by (smt (verit, ccfv_threshold) add_Suc add_Suc_right append.left_neutral eval_nat_numeral(3) + list_update_overwrite list_update_swap n_not_Suc_n nlcontents_Nil numeral_Bit0) show "ttt = 8 + (7 + nllength [] - Suc 0 + 2 * nlength 0)" using assms by simp qed text \The tapes after $t$ iterations:\ definition "tpsL t \ tps0 [(j + 1) := (\exec t <#> 0\\<^sub>N, 1), (j + 2) := nltape (map (\t. exec t <#> 0) [0..exec t <#> 1\\<^sub>N, 1), (j + 5) := nltape (map (\t. exec t <#> 1) [0..fst (exec t)\, (j + 7) := exec t 0, (j + 8) := exec t 1]" lemma lentpsL: "length (tpsL t) = k" using jk tpsL_def by simp lemma tpsL_0_eq_tps4: "tpsL 0 = tps4" proof - have *: "exec 0 = (0, [(\xs\, 0), (\[]\, 0)])" using start_config_def contents_def by auto show ?thesis (is "tpsL 0 = ?rhs") proof (rule nth_equalityI) show "length (tpsL 0) = length ?rhs" by (simp add: tps4_def tpsL_def) show "tpsL 0 ! i = ?rhs ! i" if "i < length (tpsL 0)" for i proof - show ?thesis proof (cases "i = j \ i = j + 1 \ i = j + 2 \ i = j + 4 \ i = j + 5 \ i = j + 6 \ i = j + 7 \ i = j + 8") case True then show ?thesis unfolding tps4_def tpsL_def using * tps0 jk by auto next case False then show ?thesis - unfolding tps4_def tpsL_def using * tps0 jk that by (smt (z3) nth_list_update_neq) + unfolding tps4_def tpsL_def using * tps0 jk that by (smt (verit) nth_list_update_neq) qed qed qed qed definition "tpsL1 t \ tps0 [j := \direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 0)\, j + 1 := (\exec t <#> 0\\<^sub>N, 1), j + 2 := nltape (map (\t. exec t <#> 0) [0..direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 1)\, j + 4 := (\exec t <#> 1\\<^sub>N, 1), j + 5 := nltape (map (\t. exec t <#> 1) [0..fst (exec (Suc t))\, j + 7 := exec (Suc t) 0, j + 8 := exec (Suc t) 1]" lemma lentpsL1: "length (tpsL1 t) = k" using jk tpsL1_def by (simp only: length_list_update) lemma tmL1 [transforms_intros]: assumes "fst (exec t) < length M" shows "transforms tmL1 (tpsL t) 1 (tpsL1 t)" unfolding tmL1_def proof (tform tps: M xs jk assms) show "j + 9 \ length (tpsL t)" using tpsL_def jk by (simp only: length_list_update) show "tpsL t ! j = \\\" using tpsL_def tps0 by (simp only: nth_list_update_eq nth_list_update_neq) show "tpsL t ! (j + 3) = \\\" using tpsL_def tps0 by (simp only: nth_list_update_eq nth_list_update_neq) show "tpsL t ! (j + 6) = \fst (exec t)\" using tpsL_def tps0 jk by (simp only: length_list_update nth_list_update_eq nth_list_update_neq) show "tpsL t ! (j + 7) = exec t 0" using tpsL_def tps0 jk by (simp only: length_list_update nth_list_update_eq nth_list_update_neq) show " tpsL t ! (j + 8) = exec t 1" using tpsL_def tps0 jk by (simp only: length_list_update nth_list_update_eq nth_list_update_neq One_nat_def) show "tpsL1 t = (tpsL t) [j := \direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 0)\, j + 3 := \direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 1)\, j + 6 := \fst (exec (Suc t))\, j + 7 := exec (Suc t) 0, j + 8 := exec (Suc t) 1]" unfolding tpsL1_def tpsL_def by (simp only: list_update_overwrite list_update_swap_less[of "j+7"] list_update_swap_less[of "j+6"] list_update_swap_less[of "j+3"] list_update_swap_less[of "j"]) qed definition "tpsL2 t \ tps0 [j := \direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 0)\, j + 1 := (\exec (Suc t) <#> 0\\<^sub>N, 1), j + 2 := nltape (map (\t. exec t <#> 0) [0..direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 1)\, j + 4 := (\exec t <#> 1\\<^sub>N, 1), j + 5 := nltape (map (\t. exec t <#> 1) [0..fst (exec (Suc t))\, j + 7 := exec (Suc t) 0, j + 8 := exec (Suc t) 1]" lemma lentpsL2: "length (tpsL2 t) = k" using jk tpsL2_def by (simp only: length_list_update) lemma tmL2 [transforms_intros]: assumes "fst (exec t) < length M" and "ttt = 17 + 4 * nlength thalt" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" unfolding tmL2_def proof (tform tps: M xs jk assms(1)) show "\t. exec t <#> 0 \ thalt" using max_head_pos_0 by simp show "j + 3 \ length (tpsL1 t)" using lentpsL1 jk by simp show "(0 :: nat) < 2" by simp show "tpsL1 t ! j = \direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 0)\" using tpsL1_def jk lentpsL1 by (simp only: length_list_update nth_list_update_eq nth_list_update_neq) show "tpsL1 t ! (j + 1) = (\exec t <#> 0\\<^sub>N, 1)" using tpsL1_def jk by (simp only: length_list_update nth_list_update_eq nth_list_update_neq) show "tpsL1 t ! (j + 2) = nltape (map (\t. snd (exec t) :#: 0) [0..exec (Suc t) <#> 0\\<^sub>N, 1), j + 2 := nltape (map (\t. exec t <#> 0) [0.. tps0 [j := \direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 0)\, j + 1 := (\exec (Suc t) <#> 0\\<^sub>N, 1), j + 2 := nltape (map (\t. exec t <#> 0) [0..direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 1)\, j + 4 := (\exec (Suc t) <#> 1\\<^sub>N, 1), j + 5 := nltape (map (\t. exec t <#> 1) [0..fst (exec (Suc t))\, j + 7 := exec (Suc t) 0, j + 8 := exec (Suc t) 1]" lemma lentpsL3: "length (tpsL3 t) = k" using jk tpsL3_def by (simp only: length_list_update) lemma tmL3 [transforms_intros]: assumes "fst (exec t) < length M" and "ttt = 33 + 8 * nlength thalt" shows "transforms tmL3 (tpsL t) ttt (tpsL3 t)" unfolding tmL3_def proof (tform tps: M jk assms(1)) show "\t. exec t <#> 1 \ thalt" using max_head_pos_1 . show "j + 3 + 3 \ length (tpsL2 t)" using tpsL2_def jk by (simp only: length_list_update) show "symbols_lt G xs" using xs . show "(1 :: nat) < 2" by simp show "tpsL2 t ! (j + 3) = \direction_to_symbol ((M ! fst (exec t)) (config_read (exec t)) [~] 1)\" using tpsL2_def jk lentpsL2 by (simp only: length_list_update nth_list_update_eq nth_list_update_neq) have "j + 3 + 1 = j + 4" by simp then show "tpsL2 t ! (j + 3 + 1) = (\exec t <#> 1\\<^sub>N, 1)" using tpsL2_def jk by (simp only: length_list_update nth_list_update_eq nth_list_update_neq) have "j + 3 + 2 = j + 5" by simp then show "tpsL2 t ! (j + 3 + 2) = nltape (map (\t. exec t <#> 1) [0..snd (exec (Suc t)) :#: 1\\<^sub>N, 1), j + 5 := nltape (map (\t. snd (exec t) :#: 1) [0..snd (exec (Suc t)) :#: 1\\<^sub>N, 1), j + 3 + 2 := nltape (map (\t. snd (exec t) :#: 1) [0.. tps0 [j + 1 := (\exec (Suc t) <#> 0\\<^sub>N, 1), j + 2 := nltape (map (\t. exec t <#> 0) [0..exec (Suc t) <#> 1\\<^sub>N, 1), j + 5 := nltape (map (\t. exec t <#> 1) [0..fst (exec (Suc t))\, j + 7 := exec (Suc t) 0, j + 8 := exec (Suc t) 1]" lemma lentpsL4: "length (tpsL4 t) = k" using jk tpsL4_def by (simp only: length_list_update) lemma tmL4: assumes "fst (exec t) < length M" and "ttt = 34 + 8 * nlength thalt" shows "transforms tmL4 (tpsL t) ttt (tpsL4 t)" unfolding tmL4_def proof (tform tps: jk assms(1) lentpsL3 lentpsL4 time: assms) have "tpsL4 t ! i = tpsL3 t ! i |:=| Suc 0" if "i = j \ i = j + 3 " for i proof (cases "i = j") case True then show ?thesis using tpsL3_def tpsL4_def jk lentpsL4 onesie_write tps0 by (simp only: length_list_update nth_list_update_eq nth_list_update_neq One_nat_def) next case False then have "i = j + 3" using that by simp then show ?thesis using tpsL3_def tpsL4_def jk lentpsL4 onesie_write tps0 by (simp only: length_list_update nth_list_update_eq nth_list_update_neq One_nat_def) qed then show "\ja. ja \ {j, j + 3} \ tpsL4 t ! ja = tpsL3 t ! ja |:=| 1" by simp have "tpsL4 t ! i = tpsL3 t ! i" if "i < length (tpsL4 t)" and "i \ j \ i \ j + 3" for i proof - consider "i = j" | "i = j + 1" | "i = j + 2" | "i = j + 3" | "i = j + 4" | "i = j + 5" | "i = j + 6" | "i = j + 7" | "i = j + 8" | "i < j" | "i > j + 8" by linarith then show ?thesis using tpsL3_def tpsL4_def that by (cases) (auto simp only: length_list_update nth_list_update_eq nth_list_update_neq) qed then show "\ja. ja < length (tpsL4 t) \ ja \ {j, j + 3} \ tpsL4 t ! ja = tpsL3 t ! ja" by simp qed lemma tpsL4_Suc: "tpsL4 t = tpsL (Suc t)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using lentpsL4 tpsL_def jk by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - consider "i = j" | "i = j + 1" | "i = j + 2" | "i = j + 3" | "i = j + 4" | "i = j + 5" | "i = j + 6" | "i = j + 7" | "i = j + 8" | "i < j" | "i > j + 8" by linarith then show ?thesis using tpsL4_def tpsL_def by (cases) (simp_all only: length_list_update nth_list_update_eq nth_list_update_neq) qed qed lemma tmL4': assumes "fst (exec t) < length M" and "ttt = 34 + 8 * nlength thalt" shows "transforms tmL4 (tpsL t) ttt (tpsL (Suc t))" using tpsL4_Suc tmL4 assms by simp lemma tmL: assumes "ttt = thalt * (36 + 8 * nlength thalt) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL thalt)" unfolding tmL_def proof (tform) have "tpsL t ! (j + 6) = \fst (exec t)\" for t using tpsL_def jk by (simp only: length_list_update nth_list_update_eq nth_list_update_neq) moreover have "j + 6 < length (tpsL t)" using jk tpsL_def by simp ultimately have *: "read (tpsL t) ! (j + 6) = fst (exec t)" for t using tapes_at_read'[of "j + 6" "tpsL t"] onesie_read[of "fst (exec t)"] by (simp add: lentpsL) show "\t. t < thalt \ read (tpsL t) ! (j + 6) < length M" using * thalt by simp show "\t. t < thalt \ transforms tmL4 (tpsL t) (34 + 8 * nlength thalt) (tpsL (Suc t))" using tmL4' * thalt(1) by simp show "\ read (tpsL thalt) ! (j + 6) < length M" using * thalt(2) by simp show "thalt * tosym (34 + 8 * nlength thalt) + 1 \ ttt" using assms by simp qed lemma tmL' [transforms_intros]: assumes "ttt = thalt * (36 + 8 * nlength thalt) + 1" shows "transforms tmL tps4 ttt (tpsL thalt)" using assms tmL tpsL_0_eq_tps4 by simp definition "tps5 \ tps0 [(j + 1) := (\exec thalt <#> 0\\<^sub>N, 1), (j + 2) := nltape (map (\t. exec t <#> 0) [0..exec thalt <#> 1\\<^sub>N, 1), (j + 5) := nltape (map (\t. exec t <#> 1) [0..fst (exec thalt)\, (j + 7) := exec thalt 0, (j + 8) := exec thalt 1]" lemma tm5: assumes "ttt = thalt * (36 + 8 * nlength thalt) + 15" shows "transforms tm5 tps0 ttt (tpsL thalt)" unfolding tm5_def by (tform tps: jk time: assms) lemma tm5' [transforms_intros]: assumes "ttt = thalt * (36 + 8 * nlength thalt) + 15" shows "transforms tm5 tps0 ttt tps5" using assms tm5 tps5_def tpsL_def by simp definition "tps6 \ tps0 [(j + 1) := (\exec thalt <#> 0\\<^sub>N, 1), (j + 2) := nltape (map (\t. exec t <#> 0) [0..exec thalt <#> 1\\<^sub>N, 1), (j + 5) := nltape (map (\t. exec t <#> 1) [0.. 0, (j + 8) := exec thalt 1]" lemma tm6 [transforms_intros]: assumes "ttt = thalt * (36 + 8 * nlength thalt) + 16" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def proof (tform tps: jk time: assms) show "tps6 = tps5[j + 6 := tps5 ! (j + 6) |:=| 1]" (is "?l = ?r") proof (rule nth_equalityI) show len: "length ?l = length ?r" using tps5_def tps6_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have i_less: "i < length ?r" using that len by simp consider "i = j" | "i = j + 1" | "i = j + 2" | "i = j + 3" | "i = j + 4" | "i = j + 5" | "i = j + 6" | "i = j + 7" | "i = j + 8" | "i < j" | "i > j + 8" by linarith then show ?thesis using i_less tps5_def tps6_def onesie_write tps0 by (cases) (simp_all only: length_list_update nth_list_update_eq nth_list_update_neq) qed qed qed definition "tps7 \ tps0 [(j + 1) := (\exec thalt <#> 0\\<^sub>N, 1), (j + 2) := (\map (\t. exec t <#> 0) [0..\<^sub>N\<^sub>L, 1), (j + 4) := (\exec thalt <#> 1\\<^sub>N, 1), (j + 5) := nltape (map (\t. exec t <#> 1) [0.. 0, (j + 8) := exec thalt 1]" lemma tm7 [transforms_intros]: assumes "ttt = thalt * (36 + 8 * nlength thalt) + 19 + nllength (map (\t. exec t <#> 0) [0..t. exec t <#> 0) [0..map (\t. exec t <#> 0) [0..\<^sub>N\<^sub>L, 1)" (is "_ = ?tp") by simp moreover have "tps7 = tps6[j + 2 := ?tp]" unfolding tps7_def tps6_def by (simp add: list_update_swap) ultimately show "tps7 = tps6[j + 2 := tps6 ! (j + 2) |#=| 1]" by simp qed definition "tps8 \ tps0 [(j + 1) := (\exec thalt <#> 0\\<^sub>N, 1), (j + 2) := (\map (\t. exec t <#> 0) [0..\<^sub>N\<^sub>L, 1), (j + 4) := (\exec thalt <#> 1\\<^sub>N, 1), (j + 5) := (\map (\t. exec t <#> 1) [0..\<^sub>N\<^sub>L, 1), (j + 7) := exec thalt 0, (j + 8) := exec thalt 1]" lemma tm8: assumes "ttt = thalt * (36 + 8 * nlength thalt) + 22 + nllength (map (\t. exec t <#> 0) [0..t. (exec t) <#> 1) [0..t. exec t <#> 1) [0..map (\t. exec t <#> 1) [0..\<^sub>N\<^sub>L, 1)" (is "_ = ?tp") by simp moreover have "tps8 = tps7[j + 5 := ?tp]" unfolding tps8_def tps7_def by (simp add: list_update_swap) ultimately show "tps8 = tps7[j + 5 := tps7 ! (j + 5) |#=| 1]" by simp qed lemma tm8': assumes "ttt = 27 * Suc thalt * (9 + 2 * nlength thalt)" shows "transforms tm8 tps0 ttt tps8" proof - have 0: "nllength (map (\t. exec t <#> 0) [0.. Suc (nlength thalt) * Suc thalt" using nllength_le_len_mult_max[of "map (\t. exec t <#> 0) [0..t. exec t <#> 1) [0.. Suc (nlength thalt) * Suc thalt" using nllength_le_len_mult_max[of "map (\t. exec t <#> 1) [0.. thalt * (36 + 8 * nlength thalt) + 22 + Suc (nlength thalt) * Suc thalt + Suc (nlength thalt) * Suc thalt" using 0 1 by linarith also have "... = thalt * (36 + 8 * nlength thalt) + 22 + 2 * Suc (nlength thalt) * Suc thalt" by simp also have "... \ Suc thalt * (36 + 8 * nlength thalt) + 22 + 2 * Suc (nlength thalt) * Suc thalt" by simp also have "... \ Suc thalt * (4 * (9 + 2 * nlength thalt)) + 22 + 2 * Suc (nlength thalt) * Suc thalt" by simp also have "... = 4 * Suc thalt * (9 + 2 * nlength thalt) + 22 + 2 * Suc (nlength thalt) * Suc thalt" by linarith also have "... = 4 * Suc thalt * (9 + 2 * nlength thalt) + 22 + Suc thalt * (2 + 2 * nlength thalt)" by simp also have "... \ 4 * Suc thalt * (9 + 2 * nlength thalt) + 22 + Suc thalt * (9 + 2 * nlength thalt)" proof - have "2 + 2 * nlength thalt \ 9 + 2 * nlength thalt" by simp then show ?thesis using Suc_mult_le_cancel1 add_le_cancel_left by blast qed also have "... = 5 * Suc thalt * (9 + 2 * nlength thalt) + 22" by linarith also have "... \ 5 * Suc thalt * (9 + 2 * nlength thalt) + 22 * Suc thalt * (9 + 2 * nlength thalt)" proof - have "1 \ Suc thalt * (9 + 2 * nlength thalt)" by simp then show ?thesis by linarith qed also have "... = 27 * Suc thalt * (9 + 2 * nlength thalt)" by linarith finally have "?ttt \ ttt" using assms by simp then show ?thesis using tm8 transforms_monotone by simp qed end (* context *) end (* locale turing_machine_list_headpos *) lemma transforms_tm_list_headposI [transforms_intros]: fixes G :: nat and j :: tapeidx and M :: machine fixes tps tps' :: "tape list" fixes thalt k ttt :: nat and xs :: "symbol list" assumes "turing_machine 2 G M" assumes "length tps = k" and "k \ j + 9" and "j > 0" assumes "\t\\" "tps ! (j + 1) = (\0\\<^sub>N, 0)" "tps ! (j + 2) = (\[]\\<^sub>N\<^sub>L, 0)" "tps ! (j + 3) = \\\" "tps ! (j + 4) = (\0\\<^sub>N, 0)" "tps ! (j + 5) = (\[]\\<^sub>N\<^sub>L, 0)" "tps ! (j + 6) = \\\" "tps ! (j + 7) = (\xs\, 0)" "tps ! (j + 8) = (\[]\, 0)" assumes "ttt = 27 * Suc thalt * (9 + 2 * nlength thalt)" assumes "tps' = tps [(j + 1) := (\(execute M (start_config 2 xs) thalt) <#> 0\\<^sub>N, 1), (j + 2) := (\map (\t. (execute M (start_config 2 xs) t) <#> 0) [0..\<^sub>N\<^sub>L, 1), (j + 4) := (\(execute M (start_config 2 xs) thalt) <#> 1\\<^sub>N, 1), (j + 5) := (\map (\t. (execute M (start_config 2 xs) t) <#> 1) [0..\<^sub>N\<^sub>L, 1), (j + 7) := (execute M (start_config 2 xs) thalt) 0, (j + 8) := (execute M (start_config 2 xs) thalt) 1]" shows "transforms (tm_list_headpos G M j) tps ttt tps'" proof - interpret loc: turing_machine_list_headpos . show ?thesis using assms loc.tm8' loc.tm8_eq_tm_list_headpos loc.tps8_def by metis qed section \A Turing machine for $\Psi$ formulas\ text \ CNF formulas from the $\Psi$ family of formulas feature prominently in $\Phi$. In this section we first present a Turing machine for generating arbitrary members of this family and later a specialized one for the $\Psi$ formulas that we need for $\Phi$. \ subsection \The general case\ text \ The next Turing machine generates a representation of the CNF formula $\Psi(vs, k)$. It expects $vs$ as a list of numbers on tape $j$ and the number $k$ on tape $j + 1$. A list of lists of numbers representing $\Psi(vs, k)$ is output to tape $j + 2$. The TM iterates over the elements of $vs$. In each iteration it generates a singleton clause containing the current element of $vs$ either as positive or negative literal, depending on whether $k$ is greater than zero or equal to zero. Then it decrements the number $k$. Thus the first $k$ variable indices of $vs$ will be turned into $k$ positive literals, the rest into negative ones (provided $|vs| \geq k$). \null \ definition tm_Psi :: "tapeidx \ machine" where "tm_Psi j \ WHILE [] ; \rs. rs ! j \ \ DO tm_nextract \ j (j + 3) ;; tm_times2 (j + 3) ;; IF \rs. rs ! (j + 1) \ \ THEN tm_incr (j + 3) ELSE [] ENDIF ;; tm_to_list (j + 3) ;; tm_appendl (j + 2) (j + 3) ;; tm_decr (j + 1) ;; tm_erase_cr (j + 3) DONE ;; tm_cr (j + 2) ;; tm_erase_cr j" lemma tm_Psi_tm: assumes "0 < j" and "j + 3 < k" and "G \ 6" shows "turing_machine k G (tm_Psi j)" unfolding tm_Psi_def using Nil_tm tm_nextract_tm tm_times2_tm tm_incr_tm tm_to_list_tm tm_appendl_tm tm_decr_tm tm_erase_cr_tm tm_cr_tm turing_machine_loop_turing_machine turing_machine_branch_turing_machine assms by simp text \ Two lemmas to help with the running time bound of @{const tm_Psi}: \ lemma sum_list_mono_nth: fixes xs :: "'a list" and f g :: "'a \ nat" assumes "\i g (xs ! i)" shows "sum_list (map f xs) \ sum_list (map g xs)" using assms by (metis in_set_conv_nth sum_list_mono) lemma sum_list_plus_const: fixes xs :: "'a list" and f :: "'a \ nat" and c :: nat shows "sum_list (map (\x. c + f x) xs) = c * length xs + sum_list (map f xs)" by (induction xs) simp_all locale turing_machine_Psi = fixes j :: tapeidx begin definition "tm1 \ tm_nextract \ j (j + 3)" definition "tm2 \ tm1 ;; tm_times2 (j + 3)" definition "tm23 \ IF \rs. rs ! (j + 1) \ \ THEN tm_incr (j + 3) ELSE [] ENDIF" definition "tm3 \ tm2 ;; tm23" definition "tm4 \ tm3 ;; tm_to_list (j + 3)" definition "tm5 \ tm4 ;; tm_appendl (j + 2) (j + 3)" definition "tm6 \ tm5 ;; tm_decr (j + 1)" definition "tm7 \ tm6 ;; tm_erase_cr (j + 3)" definition "tmL \ WHILE [] ; \rs. rs ! j \ \ DO tm7 DONE" definition "tm8 \ tmL ;; tm_cr (j + 2)" definition "tm9 \ tm8 ;; tm_erase_cr j" lemma tm9_eq_tm_Psi: "tm9 = tm_Psi j" unfolding tm9_def tm8_def tmL_def tm7_def tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def tm_Psi_def tm23_def by simp context fixes tps0 :: "tape list" and k kk :: nat and ns :: "nat list" assumes jk: "length tps0 = k" "0 < j" "j + 3 < k" and kk: "kk \ length ns" assumes tps0: "tps0 ! j = (\ns\\<^sub>N\<^sub>L, 1)" "tps0 ! (j + 1) = (\kk\\<^sub>N, 1)" "tps0 ! (j + 2) = (\[]\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j + 3) = (\[]\, 1)" begin definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j := nltape' ns t, j + 1 := (\kk - t\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0.. tape list" where "tpsL1 t \ tps0 [j := nltape' ns (Suc t), j + 1 := (\kk - t\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..ns ! t\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "t < length ns" and "ttt = 12 + 2 * nlength (ns ! t)" shows "transforms tm1 (tpsL t) ttt (tpsL1 t)" unfolding tm1_def proof (tform tps: assms(1) tpsL_def jk) show "tpsL t ! (j + 3) = (\0\\<^sub>N, 1)" using tpsL_def jk tps0 canrepr_0 by simp show "ttt = 12 + 2 * nlength 0 + 2 * nlength (ns ! t)" using assms(2) by simp show "tpsL1 t = (tpsL t) [j := nltape' ns (Suc t), j + 3 := (\ns ! t\\<^sub>N, 1)]" by (simp add: jk tps0 tpsL1_def tpsL_def list_update_swap numeral_3_eq_3) qed definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [j := nltape' ns (Suc t), j + 1 := (\kk - t\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..2 * ns ! t\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "t < length ns" and "ttt = 17 + 4 * nlength (ns ! t)" shows "transforms tm2 (tpsL t) ttt (tpsL2 t)" unfolding tm2_def by (tform tps: assms(1) tpsL1_def tpsL2_def jk time: assms(2)) definition tpsL3 :: "nat \ tape list" where "tpsL3 t \ tps0 [j := nltape' ns (Suc t), j + 1 := (\kk - t\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..2 * ns ! t + (if t < kk then 1 else 0)\\<^sub>N, 1)]" lemma tm23 [transforms_intros]: assumes "t < length ns" and "ttt = 7 + 2 * nlength (2 * ns ! t)" shows "transforms tm23 (tpsL2 t) ttt (tpsL3 t)" unfolding tm23_def proof (tform tps: assms(1) tpsL2_def jk time: assms(2)) show "read (tpsL2 t) ! (j + 1) \ \ \ tpsL3 t = (tpsL2 t)[j + 3 := (\Suc (2 * ns ! t)\\<^sub>N, 1)]" using tpsL2_def tpsL3_def jk read_ncontents_eq_0 by simp show "5 + 2 * nlength (2 * ns ! t) + 2 \ ttt" using assms by simp show "\ read (tpsL2 t) ! (j + 1) \ \ \ tpsL3 t = tpsL2 t" using tpsL2_def tpsL3_def jk read_ncontents_eq_0 by simp qed lemma tm3: assumes "t < length ns" and "ttt = 24 + 4 * nlength (ns ! t) + 2 * nlength (2 * ns ! t)" shows "transforms tm3 (tpsL t) ttt (tpsL3 t)" unfolding tm3_def by (tform tps: assms jk) lemma tm3' [transforms_intros]: assumes "t < length ns" and "ttt = 26 + 6 * nlength (ns ! t)" shows "transforms tm3 (tpsL t) ttt (tpsL3 t)" proof - let ?ttt = "24 + 4 * nlength (ns ! t) + 2 * nlength (2 * ns ! t)" have "nlength (2*x) \ Suc (nlength x)" for x by (metis eq_refl gr0I less_imp_le_nat nat_0_less_mult_iff nlength_0 nlength_even_le) then have "?ttt \ 24 + 4 * nlength (ns ! t) + 2 * Suc (nlength (ns ! t))" by (meson add_mono_thms_linordered_semiring(2) mult_le_mono2) then have "?ttt \ 26 + 6 * nlength (ns ! t)" by simp then show ?thesis using tm3 assms transforms_monotone by blast qed definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [j := nltape' ns (Suc t), j + 1 := (\kk - t\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..[2 * ns ! t + (if t < kk then 1 else 0)]\\<^sub>N\<^sub>L, 1)]" lemma tm4: assumes "t < length ns" and "ttt = 31 + 6 * nlength (ns ! t) + 2 * nlength (2 * ns ! t + (if t < kk then 1 else 0))" shows "transforms tm4 (tpsL t) ttt (tpsL4 t)" unfolding tm4_def proof (tform tps: assms tpsL3_def jk tps0) show "tpsL4 t = (tpsL3 t) [j + 3 := (\[2 * ns ! t + (if t < kk then 1 else 0)]\\<^sub>N\<^sub>L, 1)]" using tpsL3_def tpsL4_def jk tps0 by simp qed lemma tm4' [transforms_intros]: assumes "t < length ns" and "ttt = 33 + 8 * nlength (ns ! t)" shows "transforms tm4 (tpsL t) ttt (tpsL4 t)" proof - have "nlength (2 * ns ! t + (if t < kk then 1 else 0)) \ Suc (nlength (ns ! t))" using nlength_0_simp nlength_even_le nlength_le_n nlength_times2plus1 - by (smt (z3) add.right_neutral le_Suc_eq mult_0_right neq0_conv) + by (smt (verit) add.right_neutral le_Suc_eq mult_0_right neq0_conv) then have "31 + 6 * nlength (ns ! t) + 2 * nlength (2 * ns ! t + (if t < kk then 1 else 0)) \ ttt" using assms(2) by simp then show ?thesis using tm4 transforms_monotone assms(1) by blast qed definition tpsL5 :: "nat \ tape list" where "tpsL5 t \ tps0 [j := nltape' ns (Suc t), j + 1 := (\kk - t\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..[2 * ns ! t + (if t < kk then 1 else 0)]\\<^sub>N\<^sub>L, 1)]" lemma tm5 [transforms_intros]: assumes "t < length ns" and "ttt = 39 + 8 * nlength (ns ! t) + 2 * nllength [2 * ns ! t + (if t < kk then 1 else 0)]" shows "transforms tm5 (tpsL t) ttt (tpsL5 t)" unfolding tm5_def proof (tform tps: assms(1) tpsL4_def jk) let ?tps = "(tpsL4 t) [j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0.. tape list" where "tpsL6 t \ tps0 [j := nltape' ns (Suc t), j + 1 := (\kk - t - 1\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..[2 * ns ! t + (if t < kk then 1 else 0)]\\<^sub>N\<^sub>L, 1)]" lemma tm6: assumes "t < length ns" and "ttt = 39 + 8 * nlength (ns ! t) + 2 * nllength [2 * ns ! t + (if t < kk then 1 else 0)] + (8 + 2 * nlength (kk - t))" shows "transforms tm6 (tpsL t) ttt (tpsL6 t)" unfolding tm6_def by (tform tps: assms(1) tpsL6_def tpsL5_def jk time: assms(2)) lemma nllength_elem: "nllength [2 * ns ! t + (if t < kk then 1 else 0)] \ 2 + nlength (ns ! t)" proof - have "2 * ns ! t + (if t < kk then 1 else 0) \ 2 * ns ! t + 1" by simp then have "nlength (2 * ns ! t + (if t < kk then 1 else 0)) \ nlength (2 * ns ! t + 1)" using nlength_mono by simp then have "nlength (2 * ns ! t + (if t < kk then 1 else 0)) \ Suc (nlength (ns ! t))" using nlength_times2plus1 by fastforce then show ?thesis using nllength by simp qed lemma tm6' [transforms_intros]: assumes "t < length ns" and "ttt = 43 + 10 * nlength (ns ! t) + (8 + 2 * nlength (kk - t))" shows "transforms tm6 (tpsL t) ttt (tpsL6 t)" proof - let ?ttt = "39 + 8 * nlength (ns ! t) + 2 * nllength [2 * ns ! t + (if t < kk then 1 else 0)] + (8 + 2 * nlength (kk - t))" have "?ttt \ 39 + 8 * nlength (ns ! t) + 2 * (2 + nlength (ns ! t)) + (8 + 2 * nlength (kk - t))" using nllength_elem by (meson add_mono_thms_linordered_semiring(2) add_mono_thms_linordered_semiring(3) nat_mult_le_cancel_disj) also have "... \ 43 + 10 * nlength (ns ! t) + (8 + 2 * nlength (kk - t))" by simp finally have "?ttt \ ttt" using assms(2) by simp then show ?thesis using assms(1) tm6 transforms_monotone by blast qed definition tpsL7 :: "nat \ tape list" where "tpsL7 t \ tps0 [j := nltape' ns (Suc t), j + 1 := (\kk - Suc t\\<^sub>N, 1), j + 2 := nlltape (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..[]\, 1)]" lemma tm7: assumes "t < length ns" and "ttt = 51 + (10 * nlength (ns ! t) + 2 * nlength (kk - t)) + (7 + 2 * length (numlist [2 * ns ! t + (if t < kk then 1 else 0)]))" shows "transforms tm7 (tpsL t) ttt (tpsL7 t)" unfolding tm7_def proof (tform tps: assms(1) tpsL6_def tpsL7_def jk time: tpsL6_def jk assms(2)) let ?ns = "[2 * ns ! t + (if t < kk then 1 else 0)]" show "tpsL6 t ::: (j + 3) = \numlist ?ns\" using tpsL6_def nlcontents_def jk by simp show "proper_symbols (numlist ?ns)" using proper_symbols_numlist by simp qed lemma tm7': assumes "t < length ns" and "ttt = 62 + 14 * nllength ns" shows "transforms tm7 (tpsL t) ttt (tpsL7 t)" proof - let ?ttt = "51 + (10 * nlength (ns ! t) + 2 * nlength (kk - t)) + (7 + 2 * length (numlist [2 * ns ! t + (if t < kk then 1 else 0)]))" have "?ttt = 58 + (10 * nlength (ns ! t) + 2 * nlength (kk - t)) + 2 * length (numlist [2 * ns ! t + (if t < kk then 1 else 0)])" by simp also have "... \ 58 + (10 * nlength (ns ! t) + 2 * nlength (kk - t)) + 2 * (2 + nlength (ns ! t))" using nllength_elem nllength_def mult_le_mono2 nat_add_left_cancel_le by metis also have "... = 62 + 12 * nlength (ns ! t) + 2 * nlength (kk - t)" by simp also have "... \ 62 + 12 * nlength (ns ! t) + 2 * nlength (length ns)" using assms(1) kk nlength_mono by simp also have "... \ 62 + 12 * nllength ns + 2 * nlength (length ns)" using assms(1) member_le_nllength by simp also have "... \ 62 + 12 * nllength ns + 2 * nllength ns" using length_le_nllength nlength_le_n by (meson add_left_mono dual_order.trans mult_le_mono2) also have "... = 62 + 14 * nllength ns" by simp finally have "?ttt \ 62 + 14 * nllength ns" . then show ?thesis using assms tm7 transforms_monotone by blast qed lemma tpsL7_eq_tpsL: "tpsL7 t = tpsL (Suc t)" unfolding tpsL7_def tpsL_def using jk tps0 - by (smt (z3) Suc_eq_plus1 add_2_eq_Suc' add_cancel_left_right add_left_cancel list_update_id list_update_swap + by (smt (verit) Suc_eq_plus1 add_2_eq_Suc' add_cancel_left_right add_left_cancel list_update_id list_update_swap num.simps(8) numeral_eq_iff numeral_eq_one_iff semiring_norm(86) zero_neq_numeral) lemma tm7'' [transforms_intros]: assumes "t < length ns" and "ttt = 62 + 14 * nllength ns" shows "transforms tm7 (tpsL t) ttt (tpsL (Suc t))" using assms tpsL7_eq_tpsL tm7' by simp lemma tmL [transforms_intros]: assumes "ttt = length ns * (62 + 14 * nllength ns + 2) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length ns))" unfolding tmL_def proof (tform) let ?t = "62 + 14 * nllength ns" show "read (tpsL t) ! j \ \" if "t < length ns" for t proof - have "tpsL t ! j = nltape' ns t" using tpsL_def jk by simp then show ?thesis using nltape'_tape_read that tapes_at_read' tpsL_def jk by (metis (no_types, lifting) add_lessD1 leD length_list_update) qed have "tpsL t ! j = nltape' ns t" for t using tpsL_def jk nlcontents_def by simp then show "\ read (tpsL (length ns)) ! j \ \" using nltape'_tape_read tpsL_def jk tapes_at_read' by (metis (no_types, lifting) add_lessD1 length_list_update order_refl) show "length ns * (62 + 14 * nllength ns + 2) + 1 \ ttt" using assms by simp qed definition tps8 :: "tape list" where "tps8 \ tps0 [j := nltape' ns (length ns), j + 1 := (\0\\<^sub>N, 1), j + 2 := nlltape' (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..ns\?nss. Suc (nllength ns))" using nlllength by simp also have "... = (\i\[0..ns. Suc (nllength ns)) ?nss = map (\i. Suc (nllength (?nss ! i))) [0..ns. Suc (nllength ns)) ?nss = map (\i. Suc (nllength [2 * ns ! i + (if i < kk then 1 else 0)])) [0..i\[0.. (\i\[0.. (\i\[0..i\[0..i\[0..i. Suc (nlength (ns ! i))) [0..n. Suc (nlength n)) ns" by (rule nth_equalityI) simp_all then show ?thesis using nllength by simp qed finally have "nlllength ?nss \ 2 * length ns + nllength ns" . then have "?ttt \ Suc (length ns * (64 + 14 * nllength ns)) + Suc (Suc (Suc (2 * length ns + nllength ns)))" by simp also have "... = 4 + length ns * (64 + 14 * nllength ns) + 2 * length ns + nllength ns" by simp also have "... = 4 + length ns * (66 + 14 * nllength ns) + nllength ns" by algebra also have "... \ 4 + nllength ns * (66 + 14 * nllength ns) + nllength ns" using length_le_nllength by simp also have "... = 4 + 67 * nllength ns + 14 * nllength ns ^ 2" by algebra also have "... \ 4 + 67 * nllength ns ^ 2 + 14 * nllength ns ^ 2" using linear_le_pow by simp also have "... = 4 + 81 * nllength ns ^ 2" by simp finally have "?ttt \ 4 + 81 * nllength ns ^ 2" . then show ?thesis using tm8 assms transforms_monotone tpsL_eq_tps0 by simp qed definition tps9 :: "tape list" where "tps9 \ tps0 [j := (\[]\, 1), j + 1 := (\0\\<^sub>N, 1), j + 2 := nlltape' (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..numlist ns\" using tps8_def jk nlcontents_def by simp show "proper_symbols (numlist ns)" using proper_symbols_numlist by simp show "ttt = 4 + 81 * (nllength ns)\<^sup>2 + (tps8 :#: j + 2 * length (numlist ns) + 6)" using tps8_def jk assms nllength_def by simp qed lemma tm9': assumes "ttt = 11 + 84 * nllength ns ^ 2" shows "transforms tm9 tps0 ttt tps9" proof - have "11 + 81 * nllength ns ^ 2 + 3 * nllength ns \ 11 + 84 * nllength ns ^ 2" using linear_le_pow by simp then show ?thesis using tm9 assms transforms_monotone by simp qed end end (* locale turing_machine_Psi *) lemma transforms_tm_PsiI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k kk :: nat and ns :: "nat list" assumes "length tps = k" "0 < j" "j + 3 < k" and "kk \ length ns" assumes "tps ! j = (\ns\\<^sub>N\<^sub>L, 1)" "tps ! (j + 1) = (\kk\\<^sub>N, 1)" "tps ! (j + 2) = (\[]\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps ! (j + 3) = (\[]\, 1)" assumes "ttt = 11 + 84 * nllength ns ^ 2" assumes "tps' = tps [j := (\[]\, 1), j + 1 := (\0\\<^sub>N, 1), j + 2 := nlltape' (map (\t. [2 * ns ! t + (if t < kk then 1 else 0)]) [0..For intervals\ text \ To construct $\Phi$ we need only $\Psi$ formulas where the variable index list is an interval $\gamma_i = [iH, (i+1)H)$. In this section we devise a Turing machine that outputs $\Psi([start, start + len), \kappa)$ for arbitrary $start$, $len$, and $\kappa$. \ definition nll_Psi :: "nat \ nat \ nat \ nat list list" where "nll_Psi start len \ \ map (\i. [2 * (start + i) + (if i < \ then 1 else 0)]) [0.. = formula_n (\ [start..)" (is "?lhs = ?rhs") proof (rule nth_equalityI) show len: "length ?lhs = length ?rhs" using nll_Psi_def Psi_def formula_n_def by simp let ?psi = "\ [start.." show "?lhs ! i = ?rhs ! i" if "i < length ?lhs" for i proof - have "i < length ?psi" using that Psi_def nll_Psi_def by simp have "i < len" using that Psi_def nll_Psi_def by simp show ?thesis proof (cases "i < \") case True then have "?psi ! i = [Pos (start + i)]" using Psi_def nth_append[of "map (\s. [Pos s]) (take \ [start..s. [Pos s]) (take \ [start..) \ len * (3 + nlength (start + len))" proof - define f :: "nat \ nat list" where "f = (\i. [2 * (start + i) + (if i < \ then 1 else 0)])" let ?nss = "map f [0..) = (\ns\?nss. Suc (nllength ns))" using nlllength f_def nll_Psi_def by simp also have "... = (\i\[0..ns. Suc (nllength ns)) (f i))" by (metis (no_types, lifting) map_eq_conv map_map o_apply) also have "... = (\i\[0.. then 1 else 0)])))" using f_def by simp also have "... = (\i\[0.. then 1 else 0)))))" using nllength by simp also have "... \ (\i\[0..i. Suc (Suc (nlength (2 * (start + i) + (if i < \ then 1 else 0))))" "\i. Suc (Suc (nlength (2 * (start + i) + 1)))"] by simp also have "... \ (\i\[0..i. Suc (Suc (nlength (2 * (start + i) + 1)))" "\i. Suc (Suc (nlength (2 * (start + len))))"] by simp also have "... = len * Suc (Suc (nlength (2 * (start + len))))" using sum_list_const[of _ "[0.. len * Suc (Suc (Suc (nlength (start + len))))" using nlength_times2 by (metis add_gr_0 eq_refl mult_le_cancel1 nlength_even_le) also have "... = len * (3 + nlength (start + len))" by (simp add: Suc3_eq_add_3) finally show ?thesis . qed lemma nlllength_nll_Psi_le': assumes "start1 \ start2" shows "nlllength (nll_Psi start1 len \) \ len * (3 + nlength (start2 + len))" proof - have "nlllength (nll_Psi start1 len \) \ len * (3 + nlength (start1 + len))" using nlllength_nll_Psi_le by simp moreover have "nlength (start1 + len) \ nlength (start2 + len)" using assms nlength_mono by simp ultimately show ?thesis by (meson add_mono_thms_linordered_semiring(2) mult_le_mono2 order_trans) qed lemma H4_nlength: fixes x y H :: nat assumes "x \ y" and "H \ 3" shows "H ^ 4 * (nlength x)\<^sup>2 \ H ^ 4 * (nlength y)\<^sup>2" using assms by (simp add: nlength_mono) text \ The next Turing machine receives on tape $j$ a number $i$, on tape $j + 1$ a number $H$, and on tape $j + 2$ a number $\kappa$. It outputs $\Psi([i \cdot H, (i + 1) \cdot H), \kappa)$. \ definition tm_Psigamma :: "tapeidx \ machine" where "tm_Psigamma j \ tm_mult j (j + 1) (j + 3) ;; tm_range (j + 3) (j + 1) (j + 4) ;; tm_copyn (j + 2) (j + 5) ;; tm_Psi (j + 4) ;; tm_erase_cr (j + 3)" lemma tm_Psigamma_tm: assumes "G \ 6" and "j + 7 < k" shows "turing_machine k G (tm_Psigamma j)" unfolding tm_Psigamma_def using assms tm_mult_tm tm_range_tm tm_copyn_tm tm_Psi_tm tm_erase_cr_tm by simp locale turing_machine_Psigamma = fixes j :: tapeidx begin definition "tm1 \ tm_mult j (j + 1) (j + 3)" definition "tm2 \ tm1 ;; tm_range (j + 3) (j + 1) (j + 4)" definition "tm3 \ tm2 ;; tm_copyn (j + 2) (j + 5)" definition "tm4 \ tm3 ;; tm_Psi (j + 4)" definition "tm5 \ tm4 ;; tm_erase_cr (j + 3)" lemma tm5_eq_tm_Psigamma: "tm5 = tm_Psigamma j" using tm5_def tm4_def tm3_def tm2_def tm1_def tm_Psigamma_def by simp context fixes tps0 :: "tape list" and H k idx \ :: nat assumes jk: "length tps0 = k" "0 < j" "j + 7 < k" and H: "H \ 3" and \: "\ \ H" assumes tps0: "tps0 ! j = (\idx\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\\\\<^sub>N, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\[]\, 1)" "tps0 ! (j + 7) = (\[]\, 1)" begin definition "tps1 \ tps0 [j + 3 := (\idx * H\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 4 + 26 * (nlength idx + nlength H) ^ 2 " shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: jk tps0 tps1_def) show "tps0 ! (j + 3) = (\0\\<^sub>N, 1)" using tps0 canrepr_0 by simp show "ttt = 4 + 26 * (nlength idx + nlength H) * (nlength idx + nlength H)" using assms by algebra qed definition "tps2 \ tps0 [j + 3 := (\idx * H\\<^sub>N, 1), j + 4 := (\[idx * H..\<^sub>N\<^sub>L, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 4 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * nlength (idx * H + H))" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps0 tps1_def tps2_def jk time: assms) show "tps1 ! (j + 4) = (\[]\\<^sub>N\<^sub>L, 1)" using tps1_def tps0 jk nlcontents_Nil by simp have "j + 4 + 1 = j + 5" by simp then show "tps1 ! (j + 4 + 1) = (\0\\<^sub>N, 1)" using tps1_def tps0 jk canrepr_0 by (metis add_left_imp_eq nth_list_update_neq numeral_eq_iff semiring_norm(83) semiring_norm(90)) have "j + 4 + 2 = j + 6" by simp then show "tps1 ! (j + 4 + 2) = (\0\\<^sub>N, 1)" using tps1_def tps0 jk canrepr_0 by (metis add_left_imp_eq nth_list_update_neq num.simps(8) numeral_eq_iff) qed definition "tps3 \ tps0 [j + 3 := (\idx * H\\<^sub>N, 1), j + 4 := (\[idx * H..\<^sub>N\<^sub>L, 1), j + 5 := (\\\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 18 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * nlength (idx * H + H)) + 3 * nlength \" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps0 tps2_def tps3_def jk) show "tps2 ! (j + 5) = (\0\\<^sub>N, 1)" using tps2_def jk tps0 canrepr_0 by simp show "ttt = 4 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * nlength (idx * H + H)) + (14 + 3 * (nlength \ + nlength 0))" using assms by simp qed definition "tps4 \ tps0 [j + 3 := (\idx * H\\<^sub>N, 1), j + 4 := (\[]\, 1), j + 5 := (\0\\<^sub>N, 1), j + 6 := nlltape' (map (\t. [2 * [idx * H.. then 1 else 0)]) [0..2 + Suc H * (43 + 9 * nlength (idx * H + H)) + 3 * nlength \ + 84 * (nllength [idx * H..2" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: tps0 tps3_def tps4_def jk \ time: assms) have *: "j + 4 + 1 = j + 5" "j + 4 + 2 = j + 6" "j + 4 + 3 = j + 7" using add.assoc by simp_all have "tps3 ! (j + 5) = (\\\\<^sub>N, 1)" using tps3_def jk by simp then show "tps3 ! (j + 4 + 1) = (\\\\<^sub>N, 1)" using * by metis have "tps3 ! (j + 6) = (\[]\\<^sub>N\<^sub>L\<^sub>L, 1)" using tps3_def jk tps0 nllcontents_Nil by simp then show "tps3 ! (j + 4 + 2) = (\[]\\<^sub>N\<^sub>L\<^sub>L, 1)" using * by metis have "tps3 ! (j + 7) = (\[]\, 1)" using tps3_def jk tps0 by simp then show "tps3 ! (j + 4 + 3) = (\[]\, 1)" using * by metis have "tps4 = tps3 [j + 4 := (\[]\, 1), j + 5 := (\0\\<^sub>N, 1), j + 6 := nlltape' (map (\t. [2 * [idx * H.. then 1 else 0)]) [0..[]\, 1), j + 4 + 1 := (\0\\<^sub>N, 1), j + 4 + 2 := nlltape' (map (\t. [2 * [idx * H.. then 1 else 0)]) [0.. tps0 [j + 3 := (\idx * H\\<^sub>N, 1), j + 4 := (\[]\, 1), j + 5 := (\0\\<^sub>N, 1), j + 6 := nlltape' (map (\t. [2 * (idx * H + t) + (if t < \ then 1 else 0)]) [0..t. [2 * [idx * H.. then 1 else 0)]) [0..t. [2 * (idx * H + t) + (if t < \ then 1 else 0)]) [0..2 + Suc H * (43 + 9 * nlength (idx * H + H)) + 3 * nlength \ + 84 * (nllength [idx * H..2" shows "transforms tm4 tps0 ttt tps4'" using tm4 tps4'_eq_tps4 assms by simp definition "tps5 \ tps0 [j + 3 := (\[]\, 1), j + 4 := (\[]\, 1), j + 5 := (\0\\<^sub>N, 1), j + 6 := nlltape' (map (\t. [2 * (idx * H + t) + (if t < \ then 1 else 0)]) [0..2 + Suc H * (43 + 9 * nlength (idx * H + H)) + 3 * nlength \ + 84 * (nllength [idx * H..2 + 2 * nlength (idx * H)" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def proof (tform tps: tps0 tps4'_def tps5_def jk \ time: assms tps4'_def jk) show "proper_symbols (canrepr (idx * H))" using proper_symbols_canrepr by simp qed definition "tps5' \ tps0 [j + 6 := (\nll_Psi (idx * H) H \\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tps5'_eq_tps5: "tps5' = tps5" using tps5'_def tps5_def jk tps0 nll_Psi_def nlltape'_0 canrepr_0 by (metis list_update_id) lemma tm5': assumes "ttt = 1851 * H^4 * (nlength (Suc idx))^2" shows "transforms tm5 tps0 ttt tps5'" proof - let ?ttt = "36 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * nlength (idx * H + H)) + 3 * nlength \ + 84 * (nllength [idx * H..2 + 2 * nlength (idx * H)" have "?ttt \ 36 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * nlength (idx * H + H)) + 3 * nlength \ + 84 * (Suc (nlength (Suc idx * H)) * H)\<^sup>2 + 2 * nlength (idx * H)" using nllength_le_len_mult_max[of "[idx * H.. 36 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * nlength (idx * H + H)) + 3 * nlength H + 84 * (Suc (nlength (Suc idx * H)) * H)\<^sup>2 + 2 * nlength (idx * H)" using \ nlength_mono by simp also have "... = 36 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * nlength (Suc idx * H)) + 3 * nlength H + 84 * (Suc (nlength (Suc idx * H)) * H)\<^sup>2 + 2 * nlength (idx * H)" by (simp add: add.commute) also have "... \ 36 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 3 * nlength H + 84 * (Suc (nlength (Suc idx * H)) * H)\<^sup>2 + 2 * nlength (idx * H)" proof - have "Suc H * (43 + 9 * nlength (Suc idx * H)) \ Suc H * (43 + 9 * (nlength (Suc idx) + nlength H))" using nlength_prod by (meson add_mono_thms_linordered_semiring(2) mult_le_mono2) then show ?thesis by simp qed also have "... \ 36 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 3 * nlength H + 84 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 2 * nlength (idx * H)" using nlength_prod Suc_le_mono add_le_mono le_refl mult_le_mono power2_nat_le_eq_le by presburger also have "... \ 36 + 26 * (nlength idx + nlength H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 3 * nlength H + 84 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 2 * (nlength idx + nlength H)" using nlength_prod Suc_le_mono add_le_mono le_refl mult_le_mono power2_nat_le_eq_le by presburger also have "... \ 36 + 26 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 3 * nlength H + 84 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 2 * (nlength idx + nlength H)" proof - have "nlength idx + nlength H \ Suc (nlength (Suc idx) + nlength H)" using nlength_mono by (metis add.commute nat_add_left_cancel_le nlength_Suc_le plus_1_eq_Suc trans_le_add2) moreover have "H > 0" using H by simp ultimately have "nlength idx + nlength H \ Suc (nlength (Suc idx) + nlength H) * H" by (metis (no_types, opaque_lifting) Suc_le_eq Suc_neq_Zero mult.assoc mult.commute mult_eq_1_iff mult_le_mono nat_mult_eq_cancel_disj) then show ?thesis by simp qed also have "... = 36 + 110 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 3 * nlength H + 2 * (nlength idx + nlength H)" by simp also have "... \ 36 + 110 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 3 * (Suc (nlength (Suc idx) + nlength H) * H)^2 + 2 * (nlength idx + nlength H)" proof - have "nlength H \ Suc (nlength (Suc idx) + nlength H) * H" using H by (simp add: nlength_le_n trans_le_add1) then have "nlength H \ (Suc (nlength (Suc idx) + nlength H) * H)^2" by (meson le_refl le_trans power2_nat_le_imp_le) then show ?thesis by simp qed also have "... = 36 + 113 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 2 * (nlength idx + nlength H)" by simp also have "... \ 36 + 113 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + Suc H * (43 + 9 * (nlength (Suc idx) + nlength H)) + 2 * (Suc (nlength (Suc idx) + nlength H) * H) ^ 2" proof - have "nlength idx + nlength H \ Suc (nlength (Suc idx) + nlength H)" using nlength_mono by (simp add: le_SucI) also have "... \ Suc (nlength (Suc idx) + nlength H) * H" using H by (metis Suc_eq_plus1 le_add2 mult.commute mult_le_mono1 nat_mult_1 numeral_eq_Suc order_trans) also have "... \ (Suc (nlength (Suc idx) + nlength H) * H)^2" by (simp add: power2_eq_square) finally have "nlength idx + nlength H \ (Suc (nlength (Suc idx) + nlength H) * H)^2" . then show ?thesis by simp qed also have "... = 79 + 115 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 9 * (nlength (Suc idx) + nlength H) + (43 + 9 * (nlength (Suc idx) + nlength H)) * H" by simp also have "... = 79 + 115 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 9 * (nlength (Suc idx) + nlength H) + 43 * H + 9 * (nlength (Suc idx) + nlength H) * H" by algebra also have "... \ 79 + 115 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 9 * (Suc (nlength (Suc idx) + nlength H) * H) ^ 2 + 43 * H + 9 * (nlength (Suc idx) + nlength H) * H" proof - have "nlength (Suc idx) + nlength H \ Suc (nlength (Suc idx) + nlength H)" by simp also have "... \ Suc (nlength (Suc idx) + nlength H) * H" using H by (metis One_nat_def add_leD1 le_refl mult_le_mono mult_numeral_1_right numeral_3_eq_3 numeral_nat(1) plus_1_eq_Suc) also have "... \ (Suc (nlength (Suc idx) + nlength H) * H) ^ 2" by (simp add: power2_eq_square) finally have "nlength (Suc idx) + nlength H \ (Suc (nlength (Suc idx) + nlength H) * H) ^ 2" . then show ?thesis by simp qed also have "... = 79 + 124 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 43 * H + 9 * (nlength (Suc idx) + nlength H) * H" by simp also have "... \ 79 + 124 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 43 * H + 9 * (Suc (nlength (Suc idx) + nlength H) * H) ^ 2" proof - have "(nlength (Suc idx) + nlength H) * H \ Suc (nlength (Suc idx) + nlength H) * H" by simp then have "(nlength (Suc idx) + nlength H) * H \ (Suc (nlength (Suc idx) + nlength H) * H) ^ 2" by (metis nat_le_linear power2_nat_le_imp_le verit_la_disequality) then show ?thesis by linarith qed also have "... = 79 + 133 * (Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 + 43 * H" by simp also have "... \ 79 + 133 * (9*H^3 * (nlength (Suc idx))^2 + 4*H^4) + 43 * H" proof - let ?m = "nlength (Suc idx)" let ?l = "Suc ?m" have "(Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 = ((?l + nlength H) * H)\<^sup>2" by simp also have "... = (?l*H + nlength H*H) ^ 2" by algebra also have "... \ (?l*H + H*H) ^ 2" using nlength_le_n by simp also have "... = (?l*H)^2 + 2*?l*H^3 + H^4" by algebra also have "... \ (?l*H)^2 + 2*?l^2*H^3 + H^4" by (metis Suc_n_not_le_n add_le_mono1 mult_le_mono1 mult_le_mono2 nat_add_left_cancel_le not_less_eq_eq power2_nat_le_imp_le) also have "... = ?l^2*(H^2 + 2*H^3) + H^4" by algebra also have "... \ ?l^2*(H^3 + 2*H^3) + H^4" proof - have "H^2 \ H^3" using pow_mono by (simp add: numeral_3_eq_3 numerals(2)) then show ?thesis by simp qed also have "... = ?l^2*3*H^3 + H^4" by simp also have "... = (?m^2 + 2 * ?m + 1)*3*H^3 + H^4" - by (smt (z3) add.commute add_Suc mult_2 nat_1_add_1 one_power2 plus_1_eq_Suc power2_sum) + by (smt (verit) add.commute add_Suc mult_2 nat_1_add_1 one_power2 plus_1_eq_Suc power2_sum) also have "... \ (?m^2 + 2 * ?m^2 + 1)*3*H^3 + H^4" using linear_le_pow by simp also have "... = (3*?m^2 + 1)*3*H^3 + H^4" by simp also have "... = 9*?m^2*H^3 + 3*H^3 + H^4" by algebra also have "... \ 9*?m^2*H^3 + 3*H^4 + H^4" using pow_mono' by simp also have "... = 9*H^3 * ?m^2 + 4*H^4" by simp finally have "(Suc (nlength (Suc idx) + nlength H) * H)\<^sup>2 \ 9*H^3 * ?m^2 + 4*H^4" . then show ?thesis by simp qed also have "... = 79 + 133 * 9*H^3 * (nlength (Suc idx))^2 + 133*4*H^4 + 43 * H" by simp also have "... \ 79 + 133 * 9*H^3 * (nlength (Suc idx))^2 + 133*4*H^4 + 43 * H^4" using linear_le_pow by simp also have "... \ 79*H^4 + 133 * 9*H^3 * (nlength (Suc idx))^2 + 133*4*H^4 + 43 * H^4" using H by simp also have "... = 654 * H^4 + 1197 * H^3 * (nlength (Suc idx))^2" by simp also have "... \ 654 * H^4 + 1197 * H^4 * (nlength (Suc idx))^2" using pow_mono' by simp also have "... \ 654 * H^4 * (nlength (Suc idx))^2 + 1197 * H^4 * (nlength (Suc idx))^2" using nlength_mono nlength_1_simp by (metis add_le_mono1 le_add1 mult_le_mono2 mult_numeral_1_right numerals(1) one_le_power plus_1_eq_Suc) also have "... = 1851 * H^4 * (nlength (Suc idx))^2" by simp finally have "?ttt \ 1851 * H^4 * (nlength (Suc idx))^2" . then show ?thesis using assms tm5 transforms_monotone tps5'_eq_tps5 by simp qed end (* context *) end (* locale turing_machine_Psigamma *) lemma transforms_tm_PsigammaI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt H k idx \ :: nat assumes "length tps = k" and "0 < j" and "j + 7 < k" and "H \ 3" and "\ \ H" assumes "tps ! j = (\idx\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\\\\<^sub>N, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\[]\, 1)" "tps ! (j + 7) = (\[]\, 1)" assumes "ttt = 1851 * H^4 * (nlength (Suc idx))^2" assumes "tps' = tps [j + 6 := (\nll_Psi (idx * H) H \\\<^sub>N\<^sub>L\<^sub>L, 1)]" shows "transforms (tm_Psigamma j) tps ttt tps'" proof - interpret loc: turing_machine_Psigamma j . show ?thesis using loc.tm5' loc.tps5'_def loc.tm5_eq_tm_Psigamma assms by simp qed section \A Turing machine for $\Upsilon$ formulas\ text \ The CNF formula $\Phi_7$ is made of CNF formulas $\Upsilon(\gamma_i)$ with $\gamma_i = [i\cdot H, (i+1)\cdot H)$. In this section we build a Turing machine that outputs such CNF formulas. \ subsection \A Turing machine for singleton clauses\ text \ The $\Upsilon$ formulas, just like the $\Psi$ formulas, are conjunctions of singleton clauses. The next Turing machine outputs singleton clauses. The Turing machine has two parameters: a Boolean $incr$ and a tape index $j$. It receives a variable index on tape $j$, a CNF formula as list of lists of numbers on tape $j + 2$ and a number $H$ on tape $j + 3$. The TM appends to the formula on tape $j + 2$ a singleton clause with a positive or negative (depending on $incr$) literal derived from the variable index. It also decrements $H$ and increments the variable index, which makes it more suitable for use in a loop constructing an $\Upsilon$ formula. Given our encoding of literals, what the TM actually does is doubling the number on tape $j + 1$ and optionally (if $incr$ is true) incrementing it. \ definition tm_times2_appendl :: "bool \ tapeidx \ machine" where "tm_times2_appendl incr j \ tm_copyn j (j + 1) ;; tm_times2 (j + 1) ;; (if incr then tm_incr (j + 1) else []) ;; tm_to_list (j + 1) ;; tm_appendl (j + 2) (j + 1) ;; tm_erase_cr (j + 1) ;; tm_incr j ;; tm_decr (j + 3)" lemma tm_times2_appendl_tm: assumes "0 < j" and "j + 3 < k" and "G \ 6" shows "turing_machine k G (tm_times2_appendl incr j)" unfolding tm_times2_appendl_def using Nil_tm tm_incr_tm tm_to_list_tm tm_appendl_tm tm_decr_tm tm_erase_cr_tm tm_times2_tm assms tm_copyn_tm by simp locale turing_machine_times2_appendl = fixes j :: tapeidx begin context fixes tps0 :: "tape list" and v H k :: nat and nss :: "nat list list" and incr :: bool assumes jk: "length tps0 = k" "0 < j" "j + 3 < k" assumes tps0: "tps0 ! j = (\v\\<^sub>N, 1)" "tps0 ! (j + 1) = (\[]\, 1)" "tps0 ! (j + 2) = nlltape nss" "tps0 ! (j + 3) = (\H\\<^sub>N, 1)" begin definition "tm1 \ tm_copyn j (j + 1)" definition "tm2 \ tm1 ;; tm_times2 (j + 1)" definition "tm3 \ tm2 ;; (if incr then tm_incr (j + 1) else [])" definition "tm4 \ tm3 ;; tm_to_list (j + 1)" definition "tm5 \ tm4 ;; tm_appendl (j + 2) (j + 1)" definition "tm6 \ tm5 ;; tm_erase_cr (j + 1)" definition "tm7 \ tm6 ;; tm_incr j" definition "tm8 \ tm7 ;; tm_decr (j + 3)" lemma tm8_eq_tm_times2appendl: "tm8 \ tm_times2_appendl incr j" using tm8_def tm7_def tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def tm_times2_appendl_def by simp definition "tps1 \ tps0 [j + 1 := (\v\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 14 + 3 * nlength v" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps1_def tps0 jk) show "tps0 ! (j + 1) = (\0\\<^sub>N, 1)" using jk tps0 canrepr_0 by simp show "ttt = 14 + 3 * (nlength v + nlength 0)" using assms by simp qed definition "tps2 \ tps0 [j + 1 := (\2 * v\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 19 + 5 * nlength v" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: tps2_def tps1_def jk assms) definition "tps3 \ tps0 [j + 1 := (\2 * v + (if incr then 1 else 0)\\<^sub>N, 1)]" lemma tm3_True: assumes "ttt = 24 + 5 * nlength v + 2 * nlength (2 * v)" and incr shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps3_def tps2_def jk) let ?t = "5 + 2 * nlength (2 * v)" have "transforms (tm_incr (j + 1)) tps2 ?t tps3" by (tform tps: tps3_def tps2_def jk assms(2)) then show "transforms (if incr then tm_incr (j + 1) else []) tps2 ?t tps3" using assms(2) by simp show "ttt = 19 + 5 * nlength v + ?t" using assms by simp qed lemma tm3_False: assumes "ttt = 19 + 5 * nlength v" and "\ incr" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps3_def tps2_def jk assms) show "transforms (if incr then tm_incr (j + 1) else []) tps2 0 tps3" using transforms_Nil jk tps3_def tps2_def assms(2) by simp qed lemma tm3: assumes "ttt = 24 + 5 * nlength v + 2 * nlength (2 * v)" shows "transforms tm3 tps0 ttt tps3" using tm3_True tm3_False assms transforms_monotone by (cases incr) simp_all lemma tm3' [transforms_intros]: assumes "ttt = 26 + 7 * nlength v" shows "transforms tm3 tps0 ttt tps3" proof - have "nlength (2 * v) \ Suc (nlength v)" using nlength_times2 by simp then show ?thesis using assms tm3 transforms_monotone by simp qed definition "tps4 \ tps0 [j + 1 := (\[2 * v + (if incr then 1 else 0)]\\<^sub>N\<^sub>L, 1)]" lemma tm4: assumes "ttt = 31 + 7 * nlength v + 2 * nlength (2 * v + (if incr then 1 else 0))" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def by (tform tps: tps4_def tps3_def jk assms) lemma tm4' [transforms_intros]: assumes "ttt = 33 + 9 * nlength v" shows "transforms tm4 tps0 ttt tps4" proof - have "nlength (2 * v + (if incr then 1 else 0)) \ Suc (nlength v)" using nlength_times2 nlength_times2plus1 by simp then show ?thesis using assms tm4 transforms_monotone by simp qed definition "tps5 \ tps0 [j + 1 := (\[2 * v + (if incr then 1 else 0)]\\<^sub>N\<^sub>L, 1), j + 2 := nlltape (nss @ [[2 * v + (if incr then 1 else 0)]])]" lemma tm5 [transforms_intros]: assumes "ttt = 39 + 9 * nlength v + 2 * nllength [2 * v + (if incr then 1 else 0)]" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def proof (tform tps: tps5_def tps4_def jk tps0) show "ttt = 33 + 9 * nlength v + (7 + nlllength nss - Suc (nlllength nss) + 2 * nllength [2 * v + (if incr then 1 else 0)])" using assms by simp qed definition "tps6 \ tps0 [j + 2 := nlltape (nss @ [[2 * v + (if incr then 1 else 0)]])]" lemma tm6: assumes "ttt = 46 + 9 * nlength v + 4 * nllength [2 * v + (if incr then 1 else 0)]" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def proof (tform tps: tps6_def tps5_def jk) let ?zs = "numlist [2 * v + (if incr then 1 else 0)]" show "tps5 ::: (j + 1) = \?zs\" using jk tps5_def nlcontents_def by simp show "proper_symbols ?zs" using proper_symbols_numlist by simp show "tps6 = tps5[j + 1 := (\[]\, 1)]" using jk tps6_def tps5_def tps0 by (metis (no_types, lifting) add_left_cancel list_update_id list_update_overwrite list_update_swap numeral_eq_one_iff semiring_norm(83)) show "ttt = 39 + 9 * nlength v + 2 * nllength [2 * v + (if incr then 1 else 0)] + (tps5 :#: (j + 1) + 2 * length (numlist [2 * v + (if incr then 1 else 0)]) + 6)" using assms nllength_def tps5_def jk by simp qed lemma tm6' [transforms_intros]: assumes "ttt = 54 + 13 * nlength v" shows "transforms tm6 tps0 ttt tps6" proof - have "nlength (2 * v + (if incr then 1 else 0)) \ Suc (nlength v)" using nlength_times2 nlength_times2plus1 by simp then show ?thesis using assms tm6 transforms_monotone nllength by simp qed definition "tps7 \ tps0 [j := (\Suc v\\<^sub>N, 1), j + 2 := nlltape (nss @ [[2 * v + (if incr then 1 else 0)]])]" lemma tm7 [transforms_intros]: assumes "ttt = 59 + 15 * nlength v" shows "transforms tm7 tps0 ttt tps7" unfolding tm7_def by (tform tps: tps7_def tps6_def tps0 jk assms) definition "tps8 \ tps0 [j := (\Suc v\\<^sub>N, 1), j + 2 := nlltape (nss @ [[2 * v + (if incr then 1 else 0)]]), j + 3 := (\H - 1\\<^sub>N, 1)]" lemma tm8: assumes "ttt = 67 + 15 * nlength v + 2 * nlength H" shows "transforms tm8 tps0 ttt tps8" unfolding tm8_def by (tform tps: tps8_def tps0 tps7_def jk time: assms) end end (* locale turing_machine_times2_appendl *) lemma transforms_tm_times2_appendlI [transforms_intros]: fixes j :: tapeidx and incr :: bool fixes tps tps' :: "tape list" and ttt v H k :: nat and nss :: "nat list list" assumes "length tps = k" and "0 < j" and "j + 3 < k" assumes "tps ! j = (\v\\<^sub>N, 1)" "tps ! (j + 1) = (\[]\, 1)" "tps ! (j + 2) = nlltape nss" "tps ! (j + 3) = (\H\\<^sub>N, 1)" assumes "ttt = 67 + 15 * nlength v + 2 * nlength H" assumes "tps' = tps [j := (\Suc v\\<^sub>N, 1), j + 2 := nlltape (nss @ [[2 * v + (if incr then 1 else 0)]]), j + 3 := (\H - 1\\<^sub>N, 1)]" shows "transforms (tm_times2_appendl incr j) tps ttt tps'" proof - interpret loc: turing_machine_times2_appendl j . show ?thesis using assms loc.tm8 loc.tps8_def loc.tm8_eq_tm_times2appendl by metis qed subsection \A Turing machine for $\Upsilon(\gamma_i)$ formulas\ text \ We will not need the general $\Upsilon$ formulas, but only $\Upsilon(\gamma_i)$ for $\gamma_i = [i\cdot H, (i + 1)\cdot H)$. Represented as list of lists of numbers they look like this (for $H \geq 3$): \ definition nll_Upsilon :: "nat \ nat \ nat list list" where "nll_Upsilon idx len \ [[2 * (idx * len) + 1], [2 * (idx * len + 1) + 1]] @ map (\i. [2 * (idx * len + i)]) [3.. 3" shows "nll_Upsilon idx len = formula_n (\ [idx*len..i. [2 * (idx * len + i)]) [3..i. [2 * (idx * len + i)]) [3..i. [2 * (idx * len + i)]) ([3..i. [2 * (idx * len + i)]) (i + 1)" using False that `length nss = 2` * by simp also have "... = [2 * (idx * len + i + 1)]" by simp finally have lhs: "?lhs ! i = [2 * (idx * len + i + 1)]" . have "?ups ! i = map (\s. [Neg s]) (drop 3 [idx*len..s. [Neg s]) (drop 3 [idx*len..s. [Neg s]) (idx * len + i + 1)" using Upsilon_def False that formula_n_def len by auto finally have "?ups ! i = [Neg (idx * len + i + 1)]" . moreover have "?rhs ! i = clause_n (?ups ! i)" using Upsilon_def False that formula_n_def len by auto ultimately have "?rhs ! i = clause_n [Neg (idx * len + i + 1)]" by simp then show ?thesis using clause_n_def lhs by simp qed qed lemma nlllength_nll_Upsilon_le: assumes "len \ 3" shows "nlllength (nll_Upsilon idx len) \ len * (4 + nlength idx + nlength len)" proof - define f :: "nat \ nat list" where "f = (\i. [2 * (idx * len + i)])" let ?nss = "map f [3..ns\?nss. Suc (nllength ns))" using nlllength f_def by simp also have "... = (\i\[3..ns. Suc (nllength ns)) (f i))" by (metis (no_types, lifting) map_eq_conv map_map o_apply) also have "... = (\i\[3..i\[3.. (\i\[3..i. Suc (Suc (nlength (2 * (idx * len + i))))" "\i. Suc (Suc (nlength (2 * (Suc idx * len))))"] by simp also have "... = Suc (Suc (nlength (2 * (Suc idx * len)))) * (len - 3)" using assms sum_list_const[of _ "[3.. Suc (Suc (Suc (nlength (Suc idx * len)))) * (len - 3)" using nlength_times2 Suc_le_mono mult_le_mono1 by presburger also have "... = (len - 3) * (3 + nlength (Suc idx * len))" by (simp add: Suc3_eq_add_3) finally have *: "nlllength ?nss \ (len - 3) * (3 + nlength (Suc idx * len))" . let ?nss2 = "[[2 * (idx * len) + 1], [2 * (idx * len + 1) + 1]]" have "nlllength ?nss2 = (\ns\?nss2. Suc (nllength ns))" using nlllength by simp also have "... = Suc (nllength [2 * (idx * len) + 1]) + Suc (nllength [2 * (idx * len + 1) + 1])" by simp also have "... = Suc (Suc (nlength (2 * (idx * len) + 1))) + Suc (Suc (nlength (2 * (idx * len + 1) + 1)))" using nllength by simp also have "... \ Suc (Suc (nlength (2 * (Suc idx * len)))) + Suc (Suc (nlength (2 * (idx * len + 1) + 1)))" using nlength_mono assms by simp also have "... \ Suc (Suc (nlength (2 * (Suc idx * len)))) + Suc (Suc (nlength (2 * (Suc idx * len))))" using nlength_mono assms by simp also have "... = 2 * Suc (Suc (nlength (2 * (Suc idx * len))))" by simp also have "... \ 2 * Suc (Suc (Suc (nlength (Suc idx * len))))" using nlength_times2 by (meson Suc_le_mono mult_le_mono nle_le) also have "... = 2 * (3 + nlength (Suc idx * len))" by simp finally have **: "nlllength ?nss2 \ 2 * (3 + nlength (Suc idx * len))" . have "nll_Upsilon idx len = ?nss2 @ ?nss" using nll_Upsilon_def f_def by simp then have "nlllength (nll_Upsilon idx len) = nlllength ?nss2 + nlllength ?nss" by (metis length_append nlllength_def numlistlist_append) then have "nlllength (nll_Upsilon idx len) \ 2 * (3 + nlength (Suc idx * len)) + (len - 3) * (3 + nlength (Suc idx * len))" using * ** by simp also have "... = (2 + (len - 3)) * (3 + nlength (Suc idx * len))" by simp also have "... = (len - 1) * (3 + nlength (Suc idx * len))" using assms Nat.le_imp_diff_is_add by fastforce also have "... \ len * (3 + nlength (Suc idx * len))" by simp also have "... \ len * (3 + nlength (Suc idx) + nlength len)" using nlength_prod by (metis ab_semigroup_add_class.add_ac(1) mult_le_mono2 nat_add_left_cancel_le) also have "... \ len * (4 + nlength idx + nlength len)" using nlength_Suc by simp finally show ?thesis . qed text \ The next Turing machine outputs CNF formulas of the shape $\Upsilon(\gamma_i)$, where $\gamma_i = [i\cdot H, (i+1)\cdot H)$. It expects a number $i$ on tape $j$ and a number $H$ on tape $j + 1$. It writes a representation of the formula to tape $j + 4$. \ definition tm_Upsilongamma :: "tapeidx \ machine" where "tm_Upsilongamma j \ tm_copyn (j + 1) (j + 5) ;; tm_mult j (j + 1) (j + 2) ;; tm_times2_appendl True (j + 2) ;; tm_times2_appendl True (j + 2) ;; tm_decr (j + 5) ;; tm_incr (j + 2) ;; WHILE [] ; \rs. rs ! (j + 5) \ \ DO tm_times2_appendl False (j + 2) DONE ;; tm_erase_cr (j + 2) ;; tm_cr (j + 4)" lemma tm_Upsilongamma_tm: assumes "0 < j" and "j + 5 < k" and "G \ 6" shows "turing_machine k G (tm_Upsilongamma j)" unfolding tm_Upsilongamma_def using tm_copyn_tm Nil_tm tm_decr_tm tm_times2_appendl_tm tm_decr_tm tm_mult_tm tm_incr_tm assms turing_machine_loop_turing_machine tm_erase_cr_tm tm_cr_tm by simp locale turing_machine_Upsilongamma = fixes j :: tapeidx begin definition "tm1 \ tm_copyn (j + 1) (j + 5)" definition "tm2 \ tm1 ;; tm_mult j (j + 1) (j + 2)" definition "tm3 \ tm2 ;; tm_times2_appendl True (j + 2)" definition "tm4 \ tm3 ;; tm_times2_appendl True (j + 2)" definition "tm5 \ tm4 ;; tm_decr (j + 5)" definition "tm6 \ tm5 ;; tm_incr (j + 2)" definition "tmB \ tm_times2_appendl False (j + 2)" definition "tmL \ WHILE [] ; \rs. rs ! (j + 5) \ \ DO tmB DONE" definition "tm7 \ tm6 ;; tmL" definition "tm8 \ tm7 ;; tm_erase_cr (j + 2)" definition "tm9 \ tm8 ;; tm_cr (j + 4)" lemma tm9_eq_tm_Upsilongamma: "tm9 = tm_Upsilongamma j" using tm9_def tm8_def tm7_def tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def tmB_def tmL_def tm_Upsilongamma_def by simp context fixes tps0 :: "tape list" and idx H k :: nat assumes jk: "length tps0 = k" "0 < j" "j + 5 < k" and H: "H \ 3" assumes tps0: "tps0 ! j = (\idx\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\[]\, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" begin definition "tps1 \ tps0 [j + 5 := (\H\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 14 + 3 * nlength H" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps1_def tps0 jk) show "tps0 ! (j + 5) = (\0\\<^sub>N, 1)" using jk tps0 canrepr_0 by simp show "ttt = 14 + 3 * (nlength H + nlength 0)" using assms by simp qed definition "tps2 \ tps0 [j + 5 := (\H\\<^sub>N, 1), j + 2 := (\idx * H\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 18 + 3 * nlength H + 26 * (nlength idx + nlength H)^2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps2_def tps1_def jk tps0) show "tps1 ! (j + 2) = (\0\\<^sub>N, 1)" using tps1_def jk canrepr_1 tps0 by (metis add_left_imp_eq canrepr_0 nth_list_update_neq' numeral_eq_iff semiring_norm(89)) show "ttt = 14 + 3 * nlength H + (4 + 26 * (nlength idx + nlength H) * (nlength idx + nlength H))" using assms by algebra qed definition "tps3 \ tps0 [j + 5 := (\H - 1\\<^sub>N, 1), j + 4 := nlltape ([[2 * (idx * H) + 1]]), j + 2 := (\idx * H + 1\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 85 + 5 * nlength H + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H)" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps3_def tps2_def jk tps0) have *: "j + 2 + 1 = j + 3" "j + 2 + 2 = j + 4" "j + 2 + 3 = j + 5" by simp_all show "tps2 ! (j + 2 + 1) = (\[]\, 1)" using jk tps2_def tps0 by (simp only: *) simp show "tps2 ! (j + 2 + 2) = nlltape []" using jk tps2_def tps0 nllcontents_Nil by (simp only: *) simp show "tps2 ! (j + 2 + 3) = (\H\\<^sub>N, 1)" using jk tps2_def tps0 by (simp only: *) simp show "tps3 = tps2 [j + 2 := (\Suc (idx * H)\\<^sub>N, 1), j + 2 + 2 := nlltape ([] @ [[2 * (idx * H) + (if True then 1 else 0)]]), j + 2 + 3 := (\H - 1\\<^sub>N, 1)]" unfolding tps3_def tps2_def by (simp only: *) (simp add: list_update_swap[of "Suc (Suc j)"] list_update_swap_less[of "j+4"]) show "ttt = 18 + 3 * nlength H + 26 * (nlength idx + nlength H)\<^sup>2 + (67 + 15 * nlength (idx * H) + 2 * nlength H)" using assms by simp qed definition "tps4 \ tps0 [j + 5 := (\H - 2\\<^sub>N, 1), j + 4 := nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]]), j + 2 := (\idx * H + 2\\<^sub>N, 1)]" lemma tm4 [transforms_intros]: assumes "ttt = 152 + 5 * nlength H + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1)" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: tps4_def tps3_def jk tps0) have *: "j + 2 + 1 = j + 3" "j + 2 + 2 = j + 4" "j + 2 + 3 = j + 5" by simp_all show "tps3 ! (j + 2 + 1) = (\[]\, 1)" using jk tps3_def tps0 by (simp only: *) simp show "tps3 ! (j + 2 + 2) = nlltape [[2 * (idx * H) + 1]]" using jk tps3_def tps0 by (simp only: *) simp show "tps3 ! (j + 2 + 3) = (\H - 1\\<^sub>N, 1)" using jk tps3_def tps0 by (simp only: *) simp have 2: "2 = Suc (Suc 0)" by simp show "tps4 = tps3 [j + 2 := (\Suc (Suc (idx * H))\\<^sub>N, 1), j + 2 + 2 := nlltape ([[2 * (idx * H) + 1]] @ [[2 * Suc (idx * H) + (if True then 1 else 0)]]), j + 2 + 3 := (\H - 1 - 1\\<^sub>N, 1)]" unfolding tps4_def tps3_def by (simp only: *) (simp add: 2 list_update_swap) show "ttt = 85 + 5 * nlength H + 26 * (nlength idx + nlength H)\<^sup>2 + 15 * nlength (idx * H) + (67 + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1))" using assms by simp qed definition "tps5 \ tps0 [j + 5 := (\H - 3\\<^sub>N, 1), j + 4 := nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]]), j + 2 := (\idx * H + 2\\<^sub>N, 1)]" lemma tm5 [transforms_intros]: assumes "ttt = 160 + 5 * nlength H + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1) + 2 * nlength (H - 2)" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def proof (tform tps: tps5_def tps4_def jk tps0) show "ttt = 152 + 5 * nlength H + 26 * (nlength idx + nlength H)\<^sup>2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1) + (8 + 2 * nlength (H - 2)) " using assms by simp qed definition "tps6 \ tps0 [j + 5 := (\H - 3\\<^sub>N, 1), j + 4 := nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]]), j + 2 := (\idx * H + 3\\<^sub>N, 1)]" lemma tm6: assumes "ttt = 165 + 5 * nlength H + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1) + 2 * nlength (H - 2) + 2 * nlength (Suc (Suc (idx * H)))" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def proof (tform tps: tps6_def tps5_def jk tps0) show "tps6 = tps5[j + 2 := (\Suc (Suc (Suc (idx * H)))\\<^sub>N, 1)]" unfolding tps5_def tps6_def by (simp only: One_nat_def Suc_1 add_2_eq_Suc' add_Suc_right numeral_3_eq_3) (simp add: list_update_swap) show "ttt = 160 + 5 * nlength H + 26 * (nlength idx + nlength H)\<^sup>2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1) + 2 * nlength (H - 2) + (5 + 2 * nlength (Suc (Suc (idx * H))))" using assms by simp qed lemma tm6' [transforms_intros]: assumes "ttt = 165 + 41 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2" shows "transforms tm6 tps0 ttt tps6" proof - let ?ttt = "165 + 5 * nlength H + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1) + 2 * nlength (H - 2) + 2 * nlength (Suc (Suc (idx * H)))" have "?ttt \ 165 + 5 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (H - 1) + 2 * nlength (H - 2) + 2 * nlength (Suc (Suc (idx * H)))" using nlength_mono by simp also have "... \ 165 + 5 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (Suc idx * H) + 2 * nlength (H - 2) + 2 * nlength (Suc (Suc (idx * H)))" using nlength_mono by simp also have "... \ 165 + 5 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2 + 15 * nlength (idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc (Suc (idx * H)))" using nlength_mono by simp also have "... \ 165 + 5 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2 + 15 * nlength (Suc idx * H) + 15 * nlength (Suc (idx * H)) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc (Suc (idx * H)))" using nlength_mono by simp also have "... \ 165 + 5 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2 + 15 * nlength (Suc idx * H) + 15 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc (Suc (idx * H)))" using nlength_mono H by simp also have "... \ 165 + 5 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2 + 15 * nlength (Suc idx * H) + 15 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H)" using nlength_mono H by simp also have "... = 165 + 41 * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)^2" using nlength_mono by simp finally have "?ttt \ ttt" using assms by simp then show ?thesis using tm6 transforms_monotone by simp qed definition "tpsL t \ tps0 [j + 5 := (\H - 3 - t\\<^sub>N, 1), j + 4 := nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..<3 + t]), j + 2 := (\idx * H + 3 + t\\<^sub>N, 1)]" lemma tpsL_eq_tps6: "tpsL 0 = tps6" using tpsL_def tps6_def by simp lemma map_Suc_append: "a \ b \ map f [a..[]\, 1)" using jk tpsL_def tps0 by (simp only: *) simp let ?nss = "[[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..<3 + t]" show "tpsL t ! (j + 2 + 2) = nlltape ?nss" using jk tpsL_def by (simp only: *) simp show "tpsL t ! (j + 2 + 3) = (\H - 3 - t\\<^sub>N, 1)" using jk tpsL_def by (simp only: *) simp have "map (\i. [2 * (idx * H + i)]) [3..<3 + Suc t] = map (\i. [2 * (idx * H + i)]) [3..<3 + t] @ [[2 * (idx * H + 3 + t) + (if False then 1 else 0)]]" using map_Suc_append[of "3" "3 + t" "\i. [2 * (idx * H + i)]"] by simp then have "[[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..<3 + Suc t] = ?nss @ [[2 * (idx * H + 3 + t) + (if False then 1 else 0)]]" by simp then show "tpsL (Suc t) = (tpsL t) [j + 2 := (\Suc (idx * H + 3 + t)\\<^sub>N, 1), j + 2 + 2 := nlltape (?nss @ [[2 * (idx * H + 3 + t) + (if False then 1 else 0)]]), j + 2 + 3 := (\H - 3 - t - 1\\<^sub>N, 1)]" unfolding tpsL_def by (simp only: *) (simp add: list_update_swap[of "Suc (Suc j)"] list_update_swap_less[of "j + 4"]) show "ttt = 67 + 15 * nlength (idx * H + 3 + t) + 2 * nlength (H - 3 - t)" using assms by simp qed lemma tmB' [transforms_intros]: assumes "ttt = 67 + 15 * nlength (Suc idx * H) + 2 * nlength H" and "t < H - 3" shows "transforms tmB (tpsL t) ttt (tpsL (Suc t))" proof - let ?ttt = "67 + 15 * nlength (idx * H + 3 + t) + 2 * nlength (H - 3 - t)" have "?ttt \ 67 + 15 * nlength (idx * H + 3 + t) + 2 * nlength H" using nlength_mono by simp also have "... \ 67 + 15 * nlength (Suc idx * H) + 2 * nlength H" using assms(2) nlength_mono by simp finally have "?ttt \ ttt" using assms(1) by simp then show ?thesis using tmB transforms_monotone by blast qed lemma tmL: assumes "ttt = H * (70 + 17 * nlength (Suc idx * H))" shows "transforms tmL (tpsL 0) ttt (tpsL (H - 3))" unfolding tmL_def proof (tform) have "read (tpsL t) ! (j + 5) = \ \ H - 3 - t = 0" for t using jk tpsL_def read_ncontents_eq_0 by simp then show "\t. t < H - 3 \ read (tpsL t) ! (j + 5) \ \" and "\ read (tpsL (H - 3)) ! (j + 5) \ \" by simp_all show "(H - 3) * (67 + 15 * nlength (Suc idx * H) + 2 * nlength H + 2) + 1 \ ttt" (is "?lhs \ ttt") proof - have "?lhs = (H - 3) * (69 + 15 * nlength (Suc idx * H) + 2 * nlength H) + 1" by simp also have "... \ (H - 3) * (69 + 15 * nlength (Suc idx * H) + 2 * nlength (Suc idx * H)) + 1" using nlength_mono by simp also have "... = (H - 3) * (69 + 17 * nlength (Suc idx * H)) + 1" by simp also have "... \ H * (69 + 17 * nlength (Suc idx * H)) + 1" by simp also have "... \ H * (69 + 17 * nlength (Suc idx * H)) + H" using H by simp also have "... = H * (70 + 17 * nlength (Suc idx * H))" by algebra finally show "?lhs \ ttt" using assms by simp qed qed definition "tps7 \ tps0 [j + 5 := (\0\\<^sub>N, 1), j + 4 := nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..Suc idx * H\\<^sub>N, 1)]" lemma tpsL_eq_tps7: "tpsL (H - 3) = tps7" proof - let ?t = "H - 3" have "(\H - 3 - ?t\\<^sub>N, 1) = (\0\\<^sub>N, 1)" by simp moreover have "nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..<3 + ?t]) = nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..idx * H + 3 + ?t\\<^sub>N, 1) = (\Suc idx * H\\<^sub>N, 1)" using H by (simp add: add.commute) ultimately show ?thesis using tpsL_def tps7_def by simp qed lemma tmL' [transforms_intros]: assumes "ttt = H * (70 + 17 * nlength (Suc idx * H))" shows "transforms tmL tps6 ttt tps7" using tmL tpsL_eq_tps6 tpsL_eq_tps7 assms by simp lemma tm7 [transforms_intros]: assumes "ttt = 165 + 41 * nlength (H + idx * H) + 26 * (nlength idx + nlength H)\<^sup>2 + H * (70 + 17 * nlength (H + idx * H))" shows "transforms tm7 tps0 ttt tps7" unfolding tm7_def by (tform tps: tps7_def tpsL_def jk tps0 time: assms) definition "tps8 \ tps0 [j + 5 := (\0\\<^sub>N, 1), j + 4 := nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..[]\, 1)]" lemma tm8: assumes "ttt = 172 + 43 * nlength (H + idx * H) + 26 * (nlength idx + nlength H)\<^sup>2 + H * (70 + 17 * nlength (H + idx * H))" shows "transforms tm8 tps0 ttt tps8" unfolding tm8_def proof (tform tps: tps8_def tps7_def jk tps0 assms) show "proper_symbols (canrepr (H + idx * H))" using proper_symbols_canrepr by simp qed definition "tps8' \ tps0[j + 4 := nlltape (nll_Upsilon idx H)]" lemma tps8'_eq_tps8: "tps8' = tps8" proof - define tps where "tps = tps0 [j + 4 := nlltape ([[2 * (idx * H) + 1], [2 * (idx * H + 1) + 1]] @ map (\i. [2 * (idx * H + i)]) [3..2 + 70 * H + 17 * H * nlength (H + idx * H)" by algebra also have "... = 172 + 70 * H + (17 * H + 43) * nlength (H + idx * H) + 26 * (nlength idx + nlength H)\<^sup>2" by algebra also have "... = 172 + 70 * H + (17 * H + 43) * nlength (Suc idx * H) + 26 * (nlength idx + nlength H)\<^sup>2" by simp also have "... \ 172 + 70 * H + (17 * H + 43) * (nlength (Suc idx) + nlength H) + 26 * (nlength idx + nlength H)\<^sup>2" using nlength_prod by (meson add_le_mono mult_le_mono order_refl) also have "... \ 172 + 70 * H + (17 * H + 43) * (Suc (nlength idx) + nlength H) + 26 * (nlength idx + nlength H)\<^sup>2" using nlength_Suc add_le_mono le_refl mult_le_mono by presburger also have "... = 172 + 70 * H + (17 * H + 43) + (17 * H + 43) * (nlength idx + nlength H) + 26 * (nlength idx + nlength H)\<^sup>2" by simp also have "... = 215 + 87 * H + (17 * H + 43) * (nlength idx + nlength H) + 26 * (nlength idx + nlength H)\<^sup>2" by simp also have "... \ 215 + 87 * H + (17 * H + 43) * (nlength idx + nlength H)^2 + 26 * (nlength idx + nlength H)\<^sup>2" using linear_le_pow by simp also have "... = 215 + 87 * H + (17 * H + 69) * (nlength idx + nlength H)^2" by algebra also have "... \ 159 * H + (17 * H + 69) * (nlength idx + nlength H)^2" using H by simp also have "... \ 159 * H + 40 * H * (nlength idx + nlength H)^2" using H by simp also have "... \ 199 * H * (nlength idx + nlength H)^2" proof - have "nlength H > 0" using H nlength_0 by simp then have "nlength idx + nlength H \ 1" by linarith then show ?thesis by simp qed finally have "?ttt \ ttt" using assms by simp then show ?thesis using tps8'_eq_tps8 tm8 transforms_monotone by simp qed definition "tps9 \ tps0 [j + 4 := (\nll_Upsilon idx H\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tm9: assumes "ttt = 199 * H * (nlength idx + nlength H)^2 + Suc (Suc (Suc (nlllength (nll_Upsilon idx H))))" shows "transforms tm9 tps0 ttt tps9" unfolding tm9_def proof (tform tps: tps8'_def tps9_def jk tps0 assms) show "clean_tape (tps8' ! (j + 4))" using tps8'_def jk tps0 clean_tape_nllcontents by simp qed lemma tm9' [transforms_intros]: assumes "ttt = 205 * H * (nlength idx + nlength H)^2" shows "transforms tm9 tps0 ttt tps9" proof - let ?ttt = "199 * H * (nlength idx + nlength H)^2 + Suc (Suc (Suc (nlllength (nll_Upsilon idx H))))" have "?ttt \ 199 * H * (nlength idx + nlength H)^2 + Suc (Suc (Suc (H * (4 + nlength idx + nlength H))))" using nlllength_nll_Upsilon_le H by simp also have "... = 199 * H * (nlength idx + nlength H)^2 + 3 + H * (4 + nlength idx + nlength H)" by simp also have "... = 199 * H * (nlength idx + nlength H)^2 + 3 + 4 * H + H * (nlength idx + nlength H)" by algebra also have "... \ 199 * H * (nlength idx + nlength H)^2 + 5 * H + H * (nlength idx + nlength H)" using H by simp also have "... \ 199 * H * (nlength idx + nlength H)^2 + 5 * H + H * (nlength idx + nlength H)^2" using linear_le_pow by simp also have "... = 200 * H * (nlength idx + nlength H)^2 + 5 * H" by simp also have "... \ 205 * H * (nlength idx + nlength H)^2" proof - have "nlength H \ 1" using H nlength_0 by (metis less_one not_less not_numeral_le_zero) then show ?thesis by simp qed finally have "?ttt \ ttt" using assms by simp then show ?thesis using tm9 transforms_monotone by simp qed end (* context *) end (* locale turing_machine_Upsilongamma *) lemma transforms_tm_UpsilongammaI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt idx H k :: nat assumes "length tps = k" and "0 < j" and "j + 5 < k" and "H \ 3" assumes "tps ! j = (\idx\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\[]\, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" assumes "ttt = 205 * H * (nlength idx + nlength H)^2" assumes "tps' = tps[j + 4 := (\nll_Upsilon idx H\\<^sub>N\<^sub>L\<^sub>L, 1)]" shows "transforms (tm_Upsilongamma j) tps ttt tps'" proof - interpret loc: turing_machine_Upsilongamma j . show ?thesis using assms loc.tm9_eq_tm_Upsilongamma loc.tm9' loc.tps9_def by simp qed end diff --git a/thys/Cook_Levin/Basics.thy b/thys/Cook_Levin/Basics.thy --- a/thys/Cook_Levin/Basics.thy +++ b/thys/Cook_Levin/Basics.thy @@ -1,2776 +1,2776 @@ chapter \Introduction\ text \ The Cook-Levin theorem states that the problem \SAT{} of deciding the satisfiability of Boolean formulas in conjunctive normal form is $\NP$-complete~\cite{Cook,Levin}. This article formalizes a proof of this theorem based on the textbook \emph{Computational Complexity:\ A Modern Approach} by Arora and Barak~\cite{ccama}. \ section \Outline\ text \ We start out in Chapter~\ref{s:TM} with a definition of multi-tape Turing machines (TMs) slightly modified from Arora and Barak's definition. The remainder of the chapter is devoted to constructing ever more complex machines for arithmetic on binary numbers, evaluating polynomials, and performing basic operations on lists of numbers and even lists of lists of numbers. Specifying Turing machines and proving their correctness and running time is laborious at the best of times. We slightly alleviate the seemingly inevitable tedium of this by defining elementary reusable Turing machines and introducing ways of composing them sequentially as well as in if-then-else branches and while loops. Together with the representation of natural numbers and lists, we thus get something faintly resembling a structured programming language of sorts. In Chapter~\ref{s:TC} we introduce some basic concepts of complexity theory, such as $\mathcal{P}$, $\NP$, and polynomial-time many-one reduction. Following Arora and Barak the complexity class $\NP$ is defined via verifier Turing machines rather than nondeterministic machines, and so the deterministic TMs introduced in the previous chapter suffice for all definitions. To flesh out the chapter a little we formalize obvious proofs of $\mathcal{P} \subseteq \NP$ and the transitivity of the reducibility relation, although neither result is needed for proving the Cook-Levin theorem. Chapter~\ref{s:Sat} introduces the problem \SAT{} as a language over bit strings. Boolean formulas in conjunctive normal form (CNF) are represented as lists of clauses, each consisting of a list of literals encoded in binary numbers. The list of lists of numbers ``data type'' defined in Chapter~\ref{s:TM} will come in handy at this point. The proof of the Cook-Levin theorem has two parts: Showing that \SAT{} is in $\NP$ and showing that \SAT{} is $\NP$-hard, that is, that every language in $\NP$ can be reduced to \SAT{} in polynomial time. The first part, also proved in Chapter~\ref{s:Sat}, is fairly easy: For a satisfiable CNF formula, a satisfying assignment can be given in roughly the size of the formula, because only the variables in the formula need be assigned a truth value. Moreover whether an assignment satisfies a CNF formula can be verified easily. The hard part is showing the $\NP$-hardness of \SAT{}. The first step (Chapter~\ref{s:oblivious}) is to show that every polynomial-time computation on a multi-tape TM can be performed in polynomial time on a two-tape \emph{oblivious} TM. Oblivious means that the sequence of positions of the Turing machine's tape heads depends only on the \emph{length} of the input. Thus any language in $\NP$ has a polynomial-time two-tape oblivious verifier TM. In Chapter~\ref{s:Reducing} the proof goes on to show how the computations of such a machine can be mapped to CNF formulas such that a CNF formula is satisfiable if and only if the underlying computation was for a string in the language \SAT{} paired with a certificate. Finally in Chapter~\ref{s:Aux_TM} and Chapter~\ref{s:Red_TM} we construct a Turing machine that carries out the reduction in polynomial time. \ section \Related work\ text \ The Cook-Levin theorem has been formalized before. Gamboa and Cowles~\cite{Gamboa2004AMP} present a formalization in ACL2~\cite{acl2}. They formalize $\NP$ and reducibility in terms of Turing machines, but analyze the running time of the reduction from $\NP$-languages to \SAT{} in a different, somewhat ad-hoc, model of computation that they call ``the major weakness'' of their formalization. Employing Coq~\cite{coq}, Gäher and Kunze~\cite{Gher2021MechanisingCT} define $\NP$ and reducibility in the computational model ``call-by-value $\lambda$-calculus L'' introduced by Forster and Smolka~\cite{Forster2017WeakCL}. They show the $\NP$-completeness of \SAT{} in this framework. Turing machines appear in an intermediate problem in the chain of reductions from $\NP$ languages to \SAT{}, but are not used to show the polynomiality of the reduction. Nevertheless, this is likely the first formalization of the Cook-Levin theorem where both the complexity theoretic concepts and the proof of the polynomiality of the reduction use the same model of computation. With regards to Isabelle, Xu~et al.~\cite{Universal_Turing_Machine-AFP} provide a formalization of single-tape Turing machines with a fixed binary alphabet in the computability theory setting and construct a universal TM. While I was putting the finishing touches on this article, Dalvit and Thiemann~\cite{Multitape_To_Singletape_TM-AFP} published a formalization of (deterministic and nondeterministic) multi-tape and single-tape Turing machines and showed how to simulate the former on the latter with quadratic slowdown. Moreover, Thiemann and Schmidinger~\cite{Multiset_Ordering_NPC-AFP} prove the $\NP$-completeness of the Multiset Ordering problem, without, however, proving the polynomial-time computability of the reduction. This article uses Turing machines as model of computation for both the complexity theoretic concepts and the running time analysis of the reduction. It is thus most similar to Gäher and Kunze's work, but has a more elementary, if not brute-force, flavor to it. \ section \The core concepts\ text \ The proof of the Cook-Levin theorem awaits us in Section~\ref{s:complete} on the very last page of this article. The way there is filled with definitions of Turing machines, correctness proofs for Turing machines, and running time-bound proofs for Turing machines, all of which can easily drown out the more relevant concepts. For instance, for verifying that the theorem on the last page really is the Cook-Levin theorem, only a small fraction of this article is relevant, namely the definitions of $\NP$-completeness and of \SAT{}. Recursively breaking down these definitions yields: \begin{itemize} \item $\NP$-completeness: Section~\ref{s:TC-NP} \begin{itemize} \item languages: Section~\ref{s:TC-NP} \item $\NP$-hard: Section~\ref{s:TC-NP} \begin{itemize} \item $\NP$: Section~\ref{s:TC-NP} \begin{itemize} \item Turing machines: Section~\ref{s:tm-basic-tm} \item computing a function: Section~\ref{s:tm-basic-comp} \item pairing strings: Section~\ref{s:tm-basic-pair} \item Big-Oh, polynomial: Section~\ref{s:tm-basic-bigoh} \end{itemize} \item polynomial-time many-one reduction: Section~\ref{s:TC-NP} \end{itemize} \end{itemize} \item \SAT{}: Section~\ref{s:sat-sat-repr} \begin{itemize} \item literal, clause, CNF formula, assignment, satisfiability: Section~\ref{s:CNF} \item representing CNF formulas as strings: Section~\ref{s:sat-sat-repr} \begin{itemize} \item string: Section~\ref{s:tm-basic-tm} \item CNF formula: Section~\ref{s:CNF} \item mapping between symbols and strings: Section~\ref{s:tm-basic-comp} \item mapping between binary and quaternary alphabets: Section~\ref{s:tm-quaternary-encoding} \item lists of lists of natural numbers: Section~\ref{s:tm-numlistlist-repr} \begin{itemize} \item binary representation of natural numbers: Section~\ref{s:tm-arithmetic-binary} \item lists of natural numbers: Section~\ref{s:tm-numlist-repr} \end{itemize} \end{itemize} \end{itemize} \end{itemize} In other words the Sections~\ref{s:tm-basic}, \ref{s:tm-arithmetic-binary}, \ref{s:tm-numlist-repr}, \ref{s:tm-numlistlist-repr}, \ref{s:tm-quaternary-encoding}, \ref{s:TC-NP}, \ref{s:CNF}, and \ref{s:sat-sat-repr} cover all definitions for formalizing the statement ``\SAT{} is $\NP$-complete''. \ chapter \Turing machines\label{s:TM}\ text \ This chapter introduces Turing machines as a model of computing functions within a running-time bound. Despite being quite intuitive, Turing machines are notoriously tedious to work with. And so most of the rest of the chapter is devoted to making this a little easier by providing means of combining TMs and a library of reusable TMs for common tasks. The basic idea (Sections~\ref{s:tm-basic} and~\ref{s:tm-trans}) is to treat Turing machines as a kind of GOTO programming language. A state of a TM corresponds to a line of code executing a rather complex command that, depending on the symbols read, can write symbols, move tape heads, and jump to another state (that is, line of code). States are identified by line numbers. This makes it easy to execute TMs in sequence by concatenating two TM ``programs''. On top of the GOTO implicit in all commands, we then define IF and WHILE in the traditional way (Section~\ref{s:tm-combining}). This makes TMs more composable. The interpretation of states as line numbers deprives TMs of the ability to memorize values ``in states'', for example, the carry bit during a binary addition. In Section~\ref{s:tm-memorizing} we recover some of this flexibility. Being able to combine TMs is helpful, but we also need TMs to combine. This takes up most of the remainder of the chapter. We start with simple operations, such as moving a tape head to the next blank symbol or copying symbols between tapes (Section~\ref{s:tm-elementary}). Extending our programming language analogy for more complex TMs, we identify tapes with variables, so that a tape contains a value of a specific type, such as a number or a list of numbers. In the remaining Sections~\ref{s:tm-arithmetic} to~\ref{s:tm-wellformed} we define these ``data types'' and devise TMs for operations over them. It would be an exaggeration to say all this makes working with Turing machines easy or fun. But at least it makes TMs somewhat more feasible to use for complexity theory, as witnessed by the subsequent chapters. \ section \Basic definitions\label{s:tm-basic}\ theory Basics imports Main begin text \ While Turing machines are fairly simple, there are still a few parts to define, especially if one allows multiple tapes and an arbitrary alphabet: states, tapes (read-only or read-write), cells, tape heads, head movements, symbols, and configurations. Beyond these are more semantic aspects like executing one or many steps of a Turing machine, its running time, and what it means for a TM to ``compute a function''. Our approach at formalizing all this must look rather crude compared to Dalvit and Thiemann's~\cite{Multitape_To_Singletape_TM-AFP}, but still it does get the job done. For lack of a better place, this section also introduces a minimal version of Big-Oh, polynomials, and a pairing function for strings. \ subsection \Multi-tape Turing machines\label{s:tm-basic-tm}\ text \ Arora and Barak~\cite[p.~11]{ccama} define multi-tape Turing machines with these features: \begin{itemize} \item There are $k \geq 2$ infinite one-directional tapes, and each has one head. \item The first tape is the input tape and read-only; the other $k - 1$ tapes can be written to. \item The tape alphabet is a finite set $\Gamma$ containing at least the blank symbol $\Box$, the start symbol $\triangleright$, and the symbols \textbf{0} and \textbf{1}. \item There is a finite set $Q$ of states with start state and halting state $q_\mathit{start}, q_\mathit{halt} \in Q$. \item The behavior is described by a transition function $\delta\colon\ Q \times \Gamma^k \to Q \times \Gamma^{k-1} \times \{L, S, R\}^k$. If the TM is in a state $q$ and the symbols $g_1,\dots,g_k$ are under the $k$ tape heads and $\delta(q, (g_1, \dots, g_k)) = (q', (g'_2, \dots, g'_k), (d_1, \dots, d_k))$, then the TM writes $g'_2, \dots, g'_k$ to the writable tapes, moves the tape heads in the direction (Left, Stay, or Right) indicated by the $d_1, \dots, d_k$ and switches to state $q'$. \end{itemize} \ subsubsection \Syntax\ text \ An obvious data type for the direction a tape head can move: \ datatype direction = Left | Stay | Right text \ We simplify the definition a bit in that we identify both symbols and states with natural numbers: \begin{itemize} \item We set $\Gamma = \{0, 1, \dots, G - 1\}$ for some $G \geq 4$ and represent the symbols $\Box$, $\triangleright$, \textbf{0}, and \textbf{1} by the numbers 0, 1, 2, and~3, respectively. We represent an alphabet $\Gamma$ by its size $G$. \item We let the set of states be of the form $\{0, 1, \dots, Q\}$ for some $Q\in\nat$ and set the start state $q_\mathit{start} = 0$ and halting state $q_\mathit{halt} = Q$. \end{itemize} The last item presents a fundamental difference to the textbook definition, because it requires that Turing machines with $q_\mathit{start} = q_\mathit{halt}$ have exactly one state, whereas the textbook definition allows them arbitrarily many states. However, if $q_\mathit{start} = q_\mathit{halt}$ then the TM starts in the halting state and thus does not actually do anything. But then it does not matter if there are other states besides that one start/halting state. Our simplified definition therefore does not restrict the expressive power of TMs. It does, however, simplify composing them. \ text \ The type @{type nat} is used for symbols and for states. \ type_synonym state = nat type_synonym symbol = nat text \ It is confusing to have the numbers 2 and 3 represent the symbols \textbf{0} and \textbf{1}. The next abbreviations try to hide this somewhat. The glyphs for symbols number~4 and~5 are chosen arbitrarily. While we will encounter Turing machines with huge alphabets, only the following symbols will be used literally: \ abbreviation (input) blank_symbol :: nat ("\") where "\ \ 0" abbreviation (input) start_symbol :: nat ("\") where "\ \ 1" abbreviation (input) zero_symbol :: nat ("\") where "\ \ 2" abbreviation (input) one_symbol :: nat ("\") where "\ \ 3" abbreviation (input) bar_symbol :: nat ("\") where "\ \ 4" abbreviation (input) sharp_symbol :: nat ("\") where "\ \ 5" no_notation abs ("\_\") text \ Tapes are infinite in one direction, so each cell can be addressed by a natural number. Likewise the position of a tape head is a natural number. The contents of a tape are represented by a mapping from cell numbers to symbols. A \emph{tape} is a pair of tape contents and head position: \ type_synonym tape = "(nat \ symbol) \ nat" text \ Our formalization of Turing machines begins with a data type representing a more general concept, which we call \emph{machine}, and later adds a predicate to define which machines are \emph{Turing} machines. In this generalization the number $k$ of tapes is arbitrary, although machines with zero tapes are of little interest. Also, all tapes are writable and the alphabet is not limited, that is, $\Gamma = \nat$. The transition function becomes $ \delta\colon\ \{0, \dots, Q\} \times \nat^k \to \{0, \dots, Q\} \times \nat^k \times \{L,S,R\}^k $ or, saving us one occurrence of~$k$, $ \delta\colon\ \{0, \dots, Q\} \times \nat^k \to \{0, \dots, Q\} \times (\nat \times \{L,S,R\})^k\;. $ The transition function $\delta$ has a fixed behavior in the state $q_{halt} = Q$ (namely making the machine do nothing). Hence $\delta$ needs to be specified only for the $Q$ states $0, \dots, Q - 1$ and thus can be given as a sequence $\delta_0, \dots, \delta_{Q-1}$ where each $\delta_q$ is a function \begin{equation} \label{eq:wf} \delta_q\colon \nat^k \to \{0, \dots, Q\} \times (\nat \times \{L,S,R\})^k. \end{equation} Going one step further we allow the machine to jump to any state in $\nat$, and we will treat any state $q \geq Q$ as a halting state. The $\delta_q$ are then \begin{equation} \label{eq:proper} \delta_q\colon \nat^k \to \nat \times (\nat \times \{L,S,R\})^k. \end{equation} Finally we allow inputs and outputs of arbitrary length, turning the $\delta_q$ into \[ \delta_q\colon \nat^* \to \nat \times (\nat \times \{L,S,R\})^*. \] Such a $\delta_q$ will be called a \emph{command}, and the elements of $\nat \times \{L,S,R\}$ will be called \emph{actions}. An action consists of writing a symbol to a tape at the current tape head position and then moving the tape head. \ type_synonym action = "symbol \ direction" text \ A command maps the list of symbols read from the tapes to a follow-up state and a list of actions. It represents the machine's behavior in one state. \ type_synonym command = "symbol list \ state \ action list" text \ Machines are then simply lists of commands. The $q$-th element of the list represents the machine's behavior in state $q$. The halting state of a machine $M$ is @{term "length M"}, but there is obviously no such element in the list. \ type_synonym machine = "command list" text \ Commands in this general form are too amorphous. We call a command \emph{well-formed} for $k$ tapes and the state space $Q$ if on reading $k$ symbols it performs $k$ actions and jumps to a state in $\{0, \dots, Q\}$. A well-formed command corresponds to (\ref{eq:wf}). \ definition wf_command :: "nat \ nat \ command \ bool" where "wf_command k Q cmd \ \gs. length gs = k \ length (snd (cmd gs)) = k \ fst (cmd gs) \ Q" text \ A well-formed command is a \emph{Turing command} for $k$ tapes and alphabet $G$ if it writes only symbols from $G$ when reading symbols from $G$ and does not write to tape $0$; that is, it writes to tape $0$ the symbol it read from tape~$0$. \ definition turing_command :: "nat \ nat \ nat \ command \ bool" where "turing_command k Q G cmd \ wf_command k Q cmd \ (\gs. length gs = k \ ((\i (\i (k > 0 \ fst (snd (cmd gs) ! 0) = gs ! 0))" text \ A \emph{Turing machine} is a machine with at least two tapes and four symbols and only Turing commands. \ definition turing_machine :: "nat \ nat \ machine \ bool" where "turing_machine k G M \ k \ 2 \ G \ 4 \ (\cmd\set M. turing_command k (length M) G cmd)" subsubsection \Semantics\ text \ Next we define the semantics of machines. The state and the list of tapes make up the \emph{configuration} of a machine. The semantics are given as functions mapping configurations to follow-up configurations. \ type_synonym config = "state \ tape list" text \ We start with the semantics of a single command. An action affects a tape in the following way. For the head movements we imagine the tapes having cell~0 at the left and the cell indices growing rightward. \ fun act :: "action \ tape \ tape" where "act (w, m) tp = ((fst tp)(snd tp:=w), case m of Left \ snd tp - 1 | Stay \ snd tp | Right \ snd tp + 1)" text \ Reading symbols from one tape, from all tapes, and from configurations: \ abbreviation tape_read :: "tape \ symbol" ("|.|") where "|.| tp \ fst tp (snd tp)" definition read :: "tape list \ symbol list" where "read tps \ map tape_read tps" abbreviation config_read :: "config \ symbol list" where "config_read cfg \ read (snd cfg)" text \ The semantics of a command: \ definition sem :: "command \ config \ config" where "sem cmd cfg \ let (newstate, actions) = cmd (config_read cfg) in (newstate, map (\(a, tp). act a tp) (zip actions (snd cfg)))" text \ The semantics of one step of a machine consist in the semantics of the command corresponding to the state the machine is in. The following definition ensures that the configuration does not change when it is in a halting state. \ definition exe :: "machine \ config \ config" where "exe M cfg \ if fst cfg < length M then sem (M ! (fst cfg)) cfg else cfg" text \ Executing a machine $M$ for multiple steps: \ fun execute :: "machine \ config \ nat \ config" where "execute M cfg 0 = cfg" | "execute M cfg (Suc t) = exe M (execute M cfg t)" text \ We have defined the semantics for arbitrary machines, but most lemmas we are going to prove about @{const exe}, @{const execute}, etc.\ will require the commands to be somewhat well-behaved, more precisely to map lists of $k$ symbols to lists of $k$ actions, as shown in (\ref{eq:proper}). We will call such commands \emph{proper}. \ abbreviation proper_command :: "nat \ command \ bool" where "proper_command k cmd \ \gs. length gs = k \ length (snd (cmd gs)) = length gs" text \ Being proper is a weaker condition than being well-formed. Since @{const exe} treats the state $Q$ and the states $q > Q$ the same, we do not need the $Q$-closure property of well-formedness for most lemmas about semantics. \ text \ Next we introduce a number of abbreviations for components of a machine and aspects of its behavior. In general, symbols between bars $|\cdot|$ represent operations on tapes, inside angle brackets $<\cdot>$ operations on configurations, between colons $:\!\cdot\!:$ operations on lists of tapes, and inside brackets $[\cdot]$ operations on state/action-list pairs. As for the symbol inside the delimiters, a dot ($.$) refers to a tape symbol, a colon ($:$) to the entire tape contents, and a hash ($\#$) to a head position; an equals sign ($=$) means some component of the left-hand side is changed. An exclamation mark ($!$) accesses an element in a list on the left-hand side term. \null \ abbreviation config_length :: "config \ nat" ("||_||") where "config_length cfg \ length (snd cfg)" abbreviation tape_move_right :: "tape \ nat \ tape" (infixl "|+|" 60) where "tp |+| n \ (fst tp, snd tp + n)" abbreviation tape_move_left :: "tape \ nat \ tape" (infixl "|-|" 60) where "tp |-| n \ (fst tp, snd tp - n)" abbreviation tape_move_to :: "tape \ nat \ tape" (infixl "|#=|" 60) where "tp |#=| n \ (fst tp, n)" abbreviation tape_write :: "tape \ symbol \ tape" (infixl "|:=|" 60) where "tp |:=| h \ ((fst tp) (snd tp := h), snd tp)" abbreviation config_tape_by_no :: "config \ nat \ tape" (infix "" 90) where "cfg j \ snd cfg ! j" abbreviation config_contents_by_no :: "config \ nat \ (nat \ symbol)" (infix "<:>" 100) where "cfg <:> j \ fst (cfg j)" abbreviation config_pos_by_no :: "config \ nat \ nat" (infix "<#>" 100) where "cfg <#> j \ snd (cfg j)" abbreviation config_symbol_read :: "config \ nat \ symbol" (infix "<.>" 100) where "cfg <.> j \ (cfg <:> j) (cfg <#> j)" abbreviation config_update_state :: "config \ nat \ config" (infix "<+=>" 90) where "cfg <+=> q \ (fst cfg + q, snd cfg)" abbreviation tapes_contents_by_no :: "tape list \ nat \ (nat \ symbol)" (infix ":::" 100) where "tps ::: j \ fst (tps ! j)" abbreviation tapes_pos_by_no :: "tape list \ nat \ nat" (infix ":#:" 100) where "tps :#: j \ snd (tps ! j)" abbreviation tapes_symbol_read :: "tape list \ nat \ symbol" (infix ":.:" 100) where "tps :.: j \ (tps ::: j) (tps :#: j)" abbreviation jump_by_no :: "state \ action list \ state" ("[*] _" [90]) where "[*] sas \ fst sas" abbreviation actions_of_cmd :: "state \ action list \ action list" ("[!!] _" [100] 100) where "[!!] sas \ snd sas" abbreviation action_by_no :: "state \ action list \ nat \ action" (infix "[!]" 90) where "sas [!] j \ snd sas ! j" abbreviation write_by_no :: "state \ action list \ nat \ symbol" (infix "[.]" 90) where "sas [.] j \ fst (sas [!] j)" abbreviation direction_by_no :: "state \ action list \ nat \ direction" (infix "[~]" 100) where "sas [~] j \ snd (sas [!] j)" text \ Symbol sequences consisting of symbols from an alphabet $G$: \ abbreviation symbols_lt :: "nat \ symbol list \ bool" where "symbols_lt G rs \ \i We will frequently have to show that commands are proper or Turing commands. \ lemma turing_commandI [intro]: assumes "\gs. length gs = k \ length ([!!] cmd gs) = length gs" and "\gs. length gs = k \ (\i. i < length gs \ gs ! i < G) \ (\j. j < length gs \ cmd gs [.] j < G)" and "\gs. length gs = k \ k > 0 \ cmd gs [.] 0 = gs ! 0" and "\gs. length gs = k \ [*] (cmd gs) \ Q" shows "turing_command k Q G cmd" using assms turing_command_def wf_command_def by simp lemma turing_commandD: assumes "turing_command k Q G cmd" and "length gs = k" shows "length ([!!] cmd gs) = length gs" and "(\i. i < length gs \ gs ! i < G) \ (\j. j < length gs \ cmd gs [.] j < G)" and "k > 0 \ cmd gs [.] 0 = gs ! 0" and "\gs. length gs = k \ [*] (cmd gs) \ Q" using assms turing_command_def wf_command_def by simp_all lemma turing_command_mono: assumes "turing_command k Q G cmd" and "Q \ Q'" shows "turing_command k Q' G cmd" using turing_command_def wf_command_def assms by auto lemma proper_command_length: assumes "proper_command k cmd" and "length gs = k" shows "length ([!!] cmd gs) = length gs" using assms by simp abbreviation proper_machine :: "nat \ machine \ bool" where "proper_machine k M \ \iiii The empty Turing machine $[]$ is the one Turing machine where the start state is the halting state, that is, $q_\mathit{start} = q_\mathit{halt} = Q = 0$. It is a Turing machine for every $k \geq 2$ and $G \geq 4$: \ lemma Nil_tm: "G \ 4 \ k \ 2 \ turing_machine k G []" using turing_machine_def by simp lemma turing_machineI [intro]: assumes "k \ 2" and "G \ 4" and "\i. i < length M \ turing_command k (length M) G (M ! i)" shows "turing_machine k G M" unfolding turing_machine_def using assms by (metis in_set_conv_nth) lemma turing_machineD: assumes "turing_machine k G M" shows "k \ 2" and "G \ 4" and "\i. i < length M \ turing_command k (length M) G (M ! i)" using turing_machine_def assms by simp_all text \ A few lemmas about @{const act}, @{const read}, and @{const sem}: \null \ lemma act: "act a tp = ((fst tp)(snd tp := fst a), case snd a of Left \ snd tp - 1 | Stay \ snd tp | Right \ snd tp + 1)" by (metis act.simps prod.collapse) lemma act_Stay: "j < length tps \ act (read tps ! j, Stay) (tps ! j) = tps ! j" by (simp add: read_def) lemma act_Right: "j < length tps \ act (read tps ! j, Right) (tps ! j) = tps ! j |+| 1" by (simp add: read_def) lemma act_Left: "j < length tps \ act (read tps ! j, Left) (tps ! j) = tps ! j |-| 1" by (simp add: read_def) lemma act_Stay': "act (h, Stay) (tps ! j) = tps ! j |:=| h" by simp lemma act_Right': "act (h, Right) (tps ! j) = tps ! j |:=| h |+| 1" by simp lemma act_Left': "act (h, Left) (tps ! j) = tps ! j |:=| h |-| 1" by simp lemma act_pos_le_Suc: "snd (act a (tps ! j)) \ Suc (snd (tps ! j))" proof - obtain w m where "a = (w, m)" by fastforce then show "snd (act a (tps ! j)) \ Suc (snd (tps ! j))" using act_Left' act_Stay' act_Right' by (cases m) simp_all qed lemma act_changes_at_most_pos: assumes "i \ snd tp" shows "fst (act (h, mv) tp) i = fst tp i" by (simp add: assms) lemma act_changes_at_most_pos': assumes "i \ snd tp" shows "fst (act a tp) i = fst tp i" by (simp add: assms act) lemma read_length: "length (read tps) = length tps" using read_def by simp lemma tapes_at_read: "j < length tps \ (q, tps) <.> j = read tps ! j" unfolding read_def by simp lemma tapes_at_read': "j < length tps \ tps :.: j = read tps ! j" unfolding read_def by simp lemma read_abbrev: "j < ||cfg|| \ read (snd cfg) ! j = cfg <.> j" unfolding read_def by simp lemma sem: "sem cmd cfg = (let rs = read (snd cfg) in (fst (cmd rs), map (\(a, tp). act a tp) (zip (snd (cmd rs)) (snd cfg))))" using sem_def read_def by (metis (no_types, lifting) case_prod_beta) lemma sem': "sem cmd cfg = (fst (cmd (read (snd cfg))), map (\(a, tp). act a tp) (zip (snd (cmd (read (snd cfg)))) (snd cfg)))" using sem_def read_def by (metis (no_types, lifting) case_prod_beta) lemma sem'': "sem cmd (q, tps) = (fst (cmd (read tps)), map (\(a, tp). act a tp) (zip (snd (cmd (read tps))) tps))" using sem' by simp lemma sem_num_tapes_raw: "proper_command k cmd \ k = ||cfg|| \ k = ||sem cmd cfg||" using sem_def read_length by (simp add: case_prod_beta) lemma sem_num_tapes2: "turing_command k Q G cmd \ k = ||cfg|| \ k = ||sem cmd cfg||" using sem_num_tapes_raw turing_commandD(1) by simp corollary sem_num_tapes2': "turing_command ||cfg|| Q G cmd \ ||cfg|| = ||sem cmd cfg||" using sem_num_tapes2 by simp corollary sem_num_tapes3: "turing_command ||cfg|| Q G cmd \ ||cfg|| = ||sem cmd cfg||" by (simp add: turing_commandD(1) sem_num_tapes_raw) lemma sem_fst: assumes "cfg' = sem cmd cfg" and "rs = read (snd cfg)" shows "fst cfg' = fst (cmd rs)" using sem by (metis (no_types, lifting) assms(1) assms(2) fstI) lemma sem_snd: assumes "proper_command k cmd" and "||cfg|| = k" and "rs = read (snd cfg)" and "j < k" shows "sem cmd cfg j = act (snd (cmd rs) ! j) (snd cfg ! j)" using assms sem' read_length by simp lemma snd_semI: assumes "proper_command k cmd" and "length tps = k" and "length tps' = k" and "\j. j < k \ act (cmd (read tps) [!] j) (tps ! j) = tps' ! j" shows "snd (sem cmd (q, tps)) = snd (q', tps')" using assms sem_snd[OF assms(1)] sem_num_tapes_raw by (metis nth_equalityI snd_conv) lemma sem_snd_tm: assumes "turing_machine k G M" and "length tps = k" and "rs = read tps" and "j < k" and "q < length M" shows "sem (M ! q) (q, tps) j = act (snd ((M ! q) rs) ! j) (tps ! j)" using assms sem_snd turing_machine_def turing_commandD(1) by (metis nth_mem snd_conv) lemma semI: assumes "proper_command k cmd" and "length tps = k" and "length tps' = k" and "fst (cmd (read tps)) = q'" and "\j. j < k \ act (cmd (read tps) [!] j) (tps ! j) = tps' ! j" shows "sem cmd (q, tps) = (q', tps')" using snd_semI[OF assms(1,2,3)] assms(4,5) sem_fst by (metis prod.exhaust_sel snd_conv) text \ Commands ignore the state element of the configuration they are applied to. \ lemma sem_state_indep: assumes "snd cfg1 = snd cfg2" shows "sem cmd cfg1 = sem cmd cfg2" using sem_def assms by simp text \ A few lemmas about @{const exe} and @{const execute}: \ lemma exe_lt_length: "fst cfg < length M \ exe M cfg = sem (M ! (fst cfg)) cfg" using exe_def by simp lemma exe_ge_length: "fst cfg \ length M \ exe M cfg = cfg" using exe_def by simp lemma exe_num_tapes: assumes "turing_machine k G M" and "k = ||cfg||" shows "k = ||exe M cfg||" using assms sem_num_tapes2 turing_machine_def exe_def by (metis nth_mem) lemma exe_num_tapes_proper: assumes "proper_machine k M" and "k = ||cfg||" shows "k = ||exe M cfg||" using assms sem_num_tapes_raw turing_machine_def exe_def by metis lemma execute_num_tapes_proper: assumes "proper_machine k M" and "k = ||cfg||" shows "k = ||execute M cfg t||" using exe_num_tapes_proper assms by (induction t) simp_all lemma execute_num_tapes: assumes "turing_machine k G M" and "k = ||cfg||" shows "k = ||execute M cfg t||" using exe_num_tapes assms by (induction t) simp_all lemma execute_after_halting: assumes "fst (execute M cfg0 t) = length M" shows "execute M cfg0 (t + n) = execute M cfg0 t" by (induction n) (simp_all add: assms exe_def) lemma execute_after_halting': assumes "fst (execute M cfg0 t) \ length M" shows "execute M cfg0 (t + n) = execute M cfg0 t" by (induction n) (simp_all add: assms exe_ge_length) corollary execute_after_halting_ge: assumes "fst (execute M cfg0 t) = length M" and "t \ t'" shows "execute M cfg0 t' = execute M cfg0 t" using execute_after_halting assms le_Suc_ex by blast corollary execute_after_halting_ge': assumes "fst (execute M cfg0 t) \ length M" and "t \ t'" shows "execute M cfg0 t' = execute M cfg0 t" using execute_after_halting' assms le_Suc_ex by blast lemma execute_additive: assumes "execute M cfg1 t1 = cfg2" and "execute M cfg2 t2 = cfg3" shows "execute M cfg1 (t1 + t2) = cfg3" using assms by (induction t2 arbitrary: cfg3) simp_all lemma turing_machine_execute_states: assumes "turing_machine k G M" and "fst cfg \ length M" and "||cfg|| = k" shows "fst (execute M cfg t) \ length M" proof (induction t) case 0 then show ?case by (simp add: assms(2)) next case (Suc t) then show ?case using turing_command_def assms(1,3) exe_def execute.simps(2) execute_num_tapes sem_fst turing_machine_def wf_command_def read_length by (smt (verit, best) nth_mem) qed text \ While running times are important, usually upper bounds for them suffice. The next predicate expresses that a machine \emph{transits} from one configuration to another one in at most a certain number of steps. \ definition transits :: "machine \ config \ nat \ config \ bool" where "transits M cfg1 t cfg2 \ \t'\t. execute M cfg1 t' = cfg2" lemma transits_monotone: assumes "t \ t'" and "transits M cfg1 t cfg2" shows "transits M cfg1 t' cfg2" using assms dual_order.trans transits_def by auto lemma transits_additive: assumes "transits M cfg1 t1 cfg2" and "transits M cfg2 t2 cfg3" shows "transits M cfg1 (t1 + t2) cfg3" proof- from assms(1) obtain t1' where 1: "t1' \ t1" "execute M cfg1 t1' = cfg2" using transits_def by auto from assms(2) obtain t2' where 2: "t2' \ t2" "execute M cfg2 t2' = cfg3" using transits_def by auto then have "execute M cfg1 (t1' + t2') = cfg3" using execute_additive 1 by simp moreover have "t1' + t2' \ t1 + t2" using "1"(1) "2"(1) by simp ultimately show ?thesis using transits_def "1"(2) "2"(2) by auto qed lemma transitsI: assumes "execute M cfg1 t' = cfg2" and "t' \ t" shows "transits M cfg1 t cfg2" unfolding transits_def using assms by auto lemma execute_imp_transits: assumes "execute M cfg1 t = cfg2" shows "transits M cfg1 t cfg2" unfolding transits_def using assms by auto text \ In the vast majority of cases we are only interested in transitions from the start state to the halting state. One way to look at it is the machine \emph{transforms} a list of tapes to another list of tapes within a certain number of steps. \ definition transforms :: "machine \ tape list \ nat \ tape list \ bool" where "transforms M tps t tps' \ transits M (0, tps) t (length M, tps')" text \ The previous predicate will be the standard way in which we express the behavior of a (Turing) machine. Consider, for example, the empty machine: \ lemma transforms_Nil: "transforms [] tps 0 tps" using transforms_def transits_def by simp lemma transforms_monotone: assumes "transforms M tps t tps'" and "t \ t'" shows "transforms M tps t' tps'" using assms transforms_def transits_monotone by simp text \ Most often the tapes will have a start symbol in the first cell followed by a finite sequence of symbols. \ definition contents :: "symbol list \ (nat \ symbol)" ("\_\") where "\xs\ i \ if i = 0 then \ else if i \ length xs then xs ! (i - 1) else \" lemma contents_at_0 [simp]: "\zs\ 0 = \" using contents_def by simp lemma contents_inbounds [simp]: "i > 0 \ i \ length zs \ \zs\ i = zs ! (i - 1)" using contents_def by simp lemma contents_outofbounds [simp]: "i > length zs \ \zs\ i = \" using contents_def by simp text \ When Turing machines are used to compute functions, they are started in a specific configuration where all tapes have the format just defined and the first tape contains the input. This is called the \emph{start configuration}~\cite[p.~13]{ccama}. \ definition start_config :: "nat \ symbol list \ config" where "start_config k xs \ (0, (\xs\, 0) # replicate (k - 1) (\[]\, 0))" lemma start_config_length: "k > 0 \ ||start_config k xs|| = k" using start_config_def contents_def by simp lemma start_config1: assumes "cfg = start_config k xs" and "0 < j" and "j < k" and "i > 0" shows "(cfg <:> j) i = \" using start_config_def contents_def assms by simp lemma start_config2: assumes "cfg = start_config k xs" and "j < k" shows "(cfg <:> j) 0 = \" using start_config_def contents_def assms by (cases "0 = j") simp_all lemma start_config3: assumes "cfg = start_config k xs" and "i > 0" and "i \ length xs" shows "(cfg <:> 0) i = xs ! (i - 1)" using start_config_def contents_def assms by simp lemma start_config4: assumes "0 < j" and "j < k" shows "snd (start_config k xs) ! j = (\i. if i = 0 then \ else \, 0)" using start_config_def contents_def assms by auto lemma start_config_pos: "j < k \ start_config k zs <#> j = 0" using start_config_def by (simp add: nth_Cons') text \ We call a symbol \emph{proper} if it is neither the blank symbol nor the start symbol. \ abbreviation proper_symbols :: "symbol list \ bool" where "proper_symbols xs \ \i Suc 0" lemma proper_symbols_append: assumes "proper_symbols xs" and "proper_symbols ys" shows "proper_symbols (xs @ ys)" using assms prop_list_append by (simp add: nth_append) lemma proper_symbols_ne0: "proper_symbols xs \ \i \" by auto lemma proper_symbols_ne1: "proper_symbols xs \ \i \" by auto text \ We call the symbols \textbf{0} and \textbf{1} \emph{bit symbols}. \ abbreviation bit_symbols :: "nat list \ bool" where "bit_symbols xs \ \i \ xs ! i = \" lemma bit_symbols_append: assumes "bit_symbols xs" and "bit_symbols ys" shows "bit_symbols (xs @ ys)" using assms prop_list_append by (simp add: nth_append) subsubsection \Basic facts about Turing machines\ text \ A Turing machine with alphabet $G$ started on a symbol sequence over $G$ will only ever have symbols from $G$ on any of its tapes. \ lemma tape_alphabet: assumes "turing_machine k G M" and "symbols_lt G zs" and "j < k" shows "((execute M (start_config k zs) t) <:> j) i < G" using assms(3) proof (induction t arbitrary: i j) case 0 have "G \ 2" using turing_machine_def assms(1) by simp then show ?case using start_config_def contents_def 0 assms(2) start_config1 start_config2 - by (smt One_nat_def Suc_1 Suc_lessD Suc_pred execute.simps(1) + by (smt (verit) One_nat_def Suc_1 Suc_lessD Suc_pred execute.simps(1) fst_conv lessI nat_less_le neq0_conv nth_Cons_0 snd_conv) next case (Suc t) let ?cfg = "execute M (start_config k zs) t" have *: "execute M (start_config k zs) (Suc t) = exe M ?cfg" by simp show ?case proof (cases "fst ?cfg \ length M") case True then have "execute M (start_config k zs) (Suc t) = ?cfg" using * exe_def by simp then show ?thesis using Suc by simp next case False then have **: "execute M (start_config k zs) (Suc t) = sem (M ! (fst ?cfg)) ?cfg" using * exe_def by simp let ?rs = "config_read ?cfg" let ?cmd = "M ! (fst ?cfg)" let ?sas = "?cmd ?rs" let ?cfg' = "sem ?cmd ?cfg" have "\jj j = act (?sas [!] j) (?cfg j)" using Suc.prems 2 len read_length sem_snd turing_commandD(1) by metis then have "?cfg' <:> j = (?cfg <:> j)(?cfg <#> j := ?sas [.] j)" using act by simp then have "(?cfg' <:> j) i < G" by (simp add: len Suc sas) then show ?thesis using ** by simp qed qed corollary read_alphabet: assumes "turing_machine k G M" and "symbols_lt G zs" shows "\ih\set (config_read (execute M (start_config k zs) t)). h < G" using read_alphabet'[OF assms] by (metis in_set_conv_nth) text \ The contents of the input tape never change. \ lemma input_tape_constant: assumes "turing_machine k G M" and "k = ||cfg||" shows "execute M cfg t <:> 0 = execute M cfg 0 <:> 0" proof (induction t) case 0 then show ?case by simp next case (Suc t) let ?cfg = "execute M cfg t" have 1: "execute M cfg (Suc t) = exe M ?cfg" by simp have 2: "length (read (snd ?cfg)) = k" using execute_num_tapes assms read_length by simp have k: "k > 0" using assms(1) turing_machine_def by simp show ?case proof (cases "fst ?cfg < length M") case True then have 3: "turing_command k (length M) G (M ! fst ?cfg)" using turing_machine_def assms(1) by simp then have "(M ! fst ?cfg) (read (snd ?cfg)) [.] 0 = read (snd ?cfg) ! 0" using turing_command_def 2 k by auto then have 4: "(M ! fst ?cfg) (read (snd ?cfg)) [.] 0 = ?cfg <.> 0" using 2 k read_abbrev read_length by auto have "execute M cfg (Suc t) <:> 0 = sem (M ! fst ?cfg) ?cfg <:> 0" using True exe_def by simp also have "... = fst (act (((M ! fst ?cfg) (read (snd ?cfg))) [!] 0) (?cfg 0))" using sem_snd 2 3 k read_length turing_commandD(1) by metis also have "... = (?cfg <:> 0) ((?cfg <#> 0):=(((M ! fst ?cfg) (read (snd ?cfg))) [.] 0))" using act by simp also have "... = (?cfg <:> 0) ((?cfg <#> 0):=?cfg <.> 0)" using 4 by simp also have "... = ?cfg <:> 0" by simp finally have "execute M cfg (Suc t) <:> 0 = ?cfg <:> 0" . then show ?thesis using Suc by simp next case False then have "execute M cfg (Suc t) = ?cfg" using exe_def by simp then show ?thesis using Suc by simp qed qed text \ A head position cannot be greater than the number of steps the machine has been running. \ lemma head_pos_le_time: assumes "turing_machine k G M" and "j < k" shows "execute M (start_config k zs) t <#> j \ t" proof (induction t) case 0 have "0 < k" using assms(1) turing_machine_def by simp then have "execute M (start_config k zs) 0 <#> j = 0" using start_config_def assms(2) start_config_pos by simp then show ?case by simp next case (Suc t) have *: "execute M (start_config k zs) (Suc t) = exe M (execute M (start_config k zs) t)" (is "_ = exe M ?cfg") by simp show ?case proof (cases "fst ?cfg = length M") case True then have "execute M (start_config k zs) (Suc t) = ?cfg" using * exe_def by simp then show ?thesis using Suc by simp next case False then have less: "fst ?cfg < length M" using assms(1) turing_machine_def by (simp add: start_config_def le_neq_implies_less turing_machine_execute_states) then have "exe M ?cfg = sem (M ! (fst ?cfg)) ?cfg" using exe_def by simp moreover have "proper_command k (M ! (fst ?cfg))" using assms(1) turing_commandD(1) less turing_machine_def nth_mem by blast ultimately have "exe M ?cfg j = act (snd ((M ! (fst ?cfg)) (config_read ?cfg)) ! j) (?cfg j)" using assms(1,2) execute_num_tapes start_config_length sem_snd by auto then have "exe M ?cfg <#> j \ Suc (?cfg <#> j)" using act_pos_le_Suc assms(1,2) execute_num_tapes start_config_length by auto then show ?thesis using * Suc.IH by simp qed qed lemma head_pos_le_halting_time: assumes "turing_machine k G M" and "fst (execute M (start_config k zs) T) = length M" and "j < k" shows "execute M (start_config k zs) t <#> j \ T" using assms execute_after_halting_ge[OF assms(2)] head_pos_le_time[OF assms(1,3)] by (metis nat_le_linear order_trans) text \ A tape cannot contain non-blank symbols at a position larger than the number of steps the Turing machine has been running, except on the input tape. \ lemma blank_after_time: assumes "i > t" and "j < k" and "0 < j" and "turing_machine k G M" shows "(execute M (start_config k zs) t <:> j) i = \" using assms(1) proof (induction t) case 0 have "execute M (start_config k zs) 0 = start_config k zs" by simp then show ?case using start_config1 assms turing_machine_def by simp next case (Suc t) have "k \ 2" using assms(2,3) by simp let ?icfg = "start_config k zs" have *: "execute M ?icfg (Suc t) = exe M (execute M ?icfg t)" by simp show ?case proof (cases "fst (execute M ?icfg t) \ length M") case True then have "execute M ?icfg (Suc t) = execute M ?icfg t" using * exe_def by simp then show ?thesis using Suc by simp next case False then have "execute M ?icfg (Suc t) <:> j = sem (M ! (fst (execute M ?icfg t))) (execute M ?icfg t) <:> j" (is "_ = sem ?cmd ?cfg <:> j") using exe_lt_length * by simp also have "... = fst (map (\(a, tp). act a tp) (zip (snd (?cmd (read (snd ?cfg)))) (snd ?cfg)) ! j)" using sem' by simp also have "... = fst (act (snd (?cmd (read (snd ?cfg))) ! j) (snd ?cfg ! j))" (is "_ = fst (act ?h (snd ?cfg ! j))") proof - have "||?cfg|| = k" using assms(2) execute_num_tapes[OF assms(4)] start_config_length turing_machine_def by simp moreover have "length (snd (?cmd (read (snd ?cfg)))) = k" using assms(4) execute_num_tapes[OF assms(4)] start_config_length turing_machine_def read_length False turing_command_def wf_command_def by simp ultimately show ?thesis using assms by simp qed finally have "execute M ?icfg (Suc t) <:> j = fst (act ?h (snd ?cfg ! j))" . moreover have "i \ ?cfg <#> j" using head_pos_le_time[OF assms(4,2)] Suc Suc_lessD leD by blast ultimately have "(execute M ?icfg (Suc t) <:> j) i = fst (?cfg j) i" using act_changes_at_most_pos by (metis prod.collapse) then show ?thesis using Suc Suc_lessD by presburger qed qed subsection \Computing a function\label{s:tm-basic-comp}\ text \ Turing machines are supposed to compute functions. The functions in question map bit strings to bit strings. We model such strings as lists of Booleans and denote the bits by @{text \} and @{text \}. \ type_synonym string = "bool list" notation False ("\") and True ("\") text \ This keeps the more abstract level of computable functions separate from the level of concrete implementations as Turing machines, which can use an arbitrary alphabet. We use the term ``string'' only for bit strings, on which functions operate, and the terms ``symbol sequence'' or ``symbols'' for the things written on the tapes of Turing machines. We translate between the two levels in a straightforward way: \ abbreviation string_to_symbols :: "string \ symbol list" where "string_to_symbols x \ map (\b. if b then \ else \) x" abbreviation symbols_to_string :: "symbol list \ string" where "symbols_to_string zs \ map (\z. z = \) zs" proposition "string_to_symbols [\, \] = [\, \]" "symbols_to_string [\, \] = [\, \]" by simp_all lemma bit_symbols_to_symbols: assumes "bit_symbols zs" shows "string_to_symbols (symbols_to_string zs) = zs" using assms by (intro nth_equalityI) auto lemma symbols_to_string_to_symbols: "symbols_to_string (string_to_symbols x) = x" by (intro nth_equalityI) simp_all lemma proper_symbols_to_symbols: "proper_symbols (string_to_symbols zs)" by simp abbreviation string_to_contents :: "string \ (nat \ symbol)" where "string_to_contents x \ \i. if i = 0 then \ else if i \ length x then (if x ! (i - 1) then \ else \) else \" lemma contents_string_to_contents: "string_to_contents xs = \string_to_symbols xs\" using contents_def by auto lemma bit_symbols_to_contents: assumes "bit_symbols ns" shows "\ns\ = string_to_contents (symbols_to_string ns)" using assms bit_symbols_to_symbols contents_string_to_contents by simp text \ Definition~1.3 in the textbook~\cite{ccama} says that for a Turing machine $M$ to compute a function $f\colon\bbOI^*\to\bbOI^*$ on input $x$, ``it halts with $f(x)$ written on its output tape.'' My initial interpretation of this phrase, and the one formalized below, was that the output is written \emph{after} the start symbol $\triangleright$ in the same fashion as the input is given on the input tape. However after inspecting the Turing machine in Example~1.1, I now believe the more likely meaning is that the output \emph{overwrites} the start symbol, although Example~1.1 precedes Definition~1.3 and might not be subject to it. One advantage of the interpretation with start symbol intact is that the output tape can then be used unchanged as the input of another Turing machine, a property we exploit in Section~\ref{s:tm-composing}. Otherwise one would have to find the start cell of the output tape and either copy the contents to another tape with start symbol or shift the string to the right and restore the start symbol. One way to find the start cell is to move the tape head left while ``marking'' the cells until one reaches an already marked cell, which can only happen when the head is in the start cell, where ``moving left'' does not actually move the head. This process will take time linear in the length of the output and thus will not change the asymptotic running time of the machine. Therefore the choice of interpretation is purely one of convenience. \null \ definition halts :: "machine \ config \ bool" where "halts M cfg \ \t. fst (execute M cfg t) = length M" lemma halts_impl_le_length: assumes "halts M cfg" shows "fst (execute M cfg t) \ length M" using assms execute_after_halting_ge' halts_def by (metis linear) definition running_time :: "machine \ config \ nat" where "running_time M cfg \ LEAST t. fst (execute M cfg t) = length M" lemma running_timeD: assumes "running_time M cfg = t" and "halts M cfg" shows "fst (execute M cfg t) = length M" and "\t'. t' < t \ fst (execute M cfg t') \ length M" using assms running_time_def halts_def not_less_Least[of _ "\t. fst (execute M cfg t) = length M"] LeastI[of "\t. fst (execute M cfg t) = length M"] by auto definition halting_config :: "machine \ config \ config" where "halting_config M cfg \ execute M cfg (running_time M cfg)" abbreviation start_config_string :: "nat \ string \ config" where "start_config_string k x \ start_config k (string_to_symbols x)" text \ Another, inconsequential, difference to the textbook definition is that we designate the second tape, rather than the last tape, as the output tape. This means that the indices for the input and output tape are fixed at~0 and~1, respectively, regardless of the total number of tapes. Next is our definition of a $k$-tape Turing machine $M$ computing a function $f$ in $T$-time: \ definition computes_in_time :: "nat \ machine \ (string \ string) \ (nat \ nat) \ bool" where "computes_in_time k M f T \ \x. halts M (start_config_string k x) \ running_time M (start_config_string k x) \ T (length x) \ halting_config M (start_config_string k x) <:> 1 = string_to_contents (f x)" lemma computes_in_time_mono: assumes "computes_in_time k M f T" and "\n. T n \ T' n" shows "computes_in_time k M f T'" using assms computes_in_time_def halts_def running_time_def halting_config_def execute_after_halting_ge by (meson dual_order.trans) text \ The definition of @{const computes_in_time} can be expressed with @{const transforms} as well, which will be more convenient for us. \ lemma halting_config_execute: assumes "fst (execute M cfg t) = length M" shows "halting_config M cfg = execute M cfg t" proof- have 1: "t \ running_time M cfg" using assms running_time_def by (simp add: Least_le) then have "fst (halting_config M cfg) = length M" using assms LeastI[of "\t. fst (execute M cfg t) = length M" t] by (simp add: halting_config_def running_time_def) then show ?thesis using execute_after_halting_ge 1 halting_config_def by metis qed lemma transforms_halting_config: assumes "transforms M tps t tps'" shows "halting_config M (0, tps) = (length M, tps')" using assms transforms_def halting_config_def halting_config_execute transits_def by (metis fst_eqD) lemma computes_in_time_execute: assumes "computes_in_time k M f T" shows "execute M (start_config_string k x) (T (length x)) <:> 1 = string_to_contents (f x)" proof - let ?t = "running_time M (start_config_string k x)" let ?cfg = "start_config_string k x" have "execute M ?cfg ?t = halting_config M ?cfg" using halting_config_def by simp then have "fst (execute M ?cfg ?t) = length M" using assms computes_in_time_def running_timeD(1) by blast moreover have "?t \ T (length x)" using computes_in_time_def assms by simp ultimately have "execute M ?cfg ?t = execute M ?cfg (T (length x)) " using execute_after_halting_ge by presburger moreover have "execute M ?cfg ?t <:> 1 = string_to_contents (f x)" using computes_in_time_def halting_config_execute assms halting_config_def by simp ultimately show ?thesis by simp qed lemma transforms_running_time: assumes "transforms M tps t tps'" shows "running_time M (0, tps) \ t" using running_time_def transforms_def transits_def - by (smt Least_le[of _ t] assms execute_after_halting_ge fst_conv) + by (smt (verit) Least_le[of _ t] assms execute_after_halting_ge fst_conv) text \ This is the alternative characterization of @{const computes_in_time}: \ lemma computes_in_time_alt: "computes_in_time k M f T = (\x. \tps. tps ::: 1 = string_to_contents (f x) \ transforms M (snd (start_config_string k x)) (T (length x)) tps)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof fix x :: string let ?cfg = "start_config_string k x" assume "computes_in_time k M f T" then have 1: "halts M ?cfg" and 2: "running_time M ?cfg \ T (length x)" and 3: "halting_config M ?cfg <:> 1 = string_to_contents (f x)" using computes_in_time_def by simp_all define cfg where "cfg = halting_config M ?cfg" then have "transits M ?cfg (T (length x)) cfg" using 2 halting_config_def transits_def by auto then have "transforms M (snd ?cfg) (T (length x)) (snd cfg)" using transits_def transforms_def start_config_def by (metis (no_types, lifting) "1" cfg_def halting_config_def prod.collapse running_timeD(1) snd_conv) moreover have "snd cfg ::: 1 = string_to_contents (f x)" using cfg_def 3 by simp ultimately show "\tps. tps ::: 1 = string_to_contents (f x) \ transforms M (snd (start_config_string k x)) (T (length x)) tps" by auto qed show "?rhs \ ?lhs" unfolding computes_in_time_def proof assume rhs: ?rhs fix x :: string let ?cfg = "start_config_string k x" obtain tps where tps: "tps ::: 1 = string_to_contents (f x)" "transforms M (snd ?cfg) (T (length x)) tps" using rhs by auto then have "transits M ?cfg (T (length x)) (length M, tps)" using transforms_def start_config_def by simp then have 1: "halts M ?cfg" using halts_def transits_def by (metis fst_eqD) moreover have 2: "running_time M ?cfg \ T (length x)" using tps(2) transforms_running_time start_config_def by simp moreover have 3: "halting_config M ?cfg <:> 1 = string_to_contents (f x)" proof - have "halting_config M ?cfg = (length M, tps)" using transforms_halting_config[OF tps(2)] start_config_def by simp then show ?thesis using tps(1) by simp qed ultimately show "halts M ?cfg \ running_time M ?cfg \ T (length x) \ halting_config M ?cfg <:> 1 = string_to_contents (f x)" by simp qed qed lemma computes_in_timeD: fixes x assumes "computes_in_time k M f T" shows "\tps. tps ::: 1 = string_to_contents (f x) \ transforms M (snd (start_config k (string_to_symbols x))) (T (length x)) tps" using assms computes_in_time_alt by simp lemma computes_in_timeI [intro]: assumes "\x. \tps. tps ::: 1 = string_to_contents (f x) \ transforms M (snd (start_config k (string_to_symbols x))) (T (length x)) tps" shows "computes_in_time k M f T" using assms computes_in_time_alt by simp text \ As an example, the function mapping every string to the empty string is computable within any time bound by the empty Turing machine. \ lemma computes_Nil_empty: assumes "k \ 2" shows "computes_in_time k [] (\x. []) T" proof fix x :: string let ?tps = "snd (start_config_string k x)" let ?f = "\x. []" have "?tps ::: 1 = string_to_contents (?f x)" using start_config4 assms by auto moreover have "transforms [] ?tps (T (length x)) ?tps" using transforms_Nil transforms_monotone by blast ultimately show "\tps. tps ::: 1 = string_to_contents (?f x) \ transforms [] ?tps (T (length x)) tps" by auto qed subsection \Pairing strings\label{s:tm-basic-pair}\ text \ In order to define the computability of functions with two arguments, we need a way to encode a pair of strings as one string. The idea is to write the two strings with a separator, for example, $\bbbO\bbbI\bbbO\bbbO\#\bbbI\bbbI\bbbI\bbbO$ and then encode every symbol $\bbbO, \bbbI, \#$ by two bits from $\bbOI$. We slightly deviate from Arora and Barak's encoding~\cite[p.~2]{ccama} and map $\bbbO$ to $\bbbO\bbbO$, $\bbbI$ to $\bbbO\bbbI$, and \# to $\bbbI\bbbI$, the idea being that the first bit signals whether the second bit is to be taken literally or as a special character. Our example turns into $\bbbO\bbbO\bbbO\bbbI\bbbO\bbbO\bbbO\bbbO\bbbI\bbbI\bbbO\bbbI\bbbO\bbbI\bbbO\bbbI\bbbO\bbbO$. \null \ abbreviation bitenc :: "string \ string" where "bitenc x \ concat (map (\h. [\, h]) x)" definition string_pair :: "string \ string \ string" ("\_, _\") where "\x, y\ \ bitenc x @ [\, \] @ bitenc y" text \ Our example: \ proposition "\[\, \, \, \], [\, \, \, \]\ = [\, \, \, \, \, \, \, \, \, \, \, \, \, \, \, \, \, \]" using string_pair_def by simp lemma length_string_pair: "length \x, y\ = 2 * length x + 2 * length y + 2" proof - have "length (concat (map (\h. [\, h]) z)) = 2 * length z" for z by (induction z) simp_all then show ?thesis using string_pair_def by simp qed lemma length_bitenc: "length (bitenc z) = 2 * length z" by (induction z) simp_all lemma bitenc_nth: assumes "i < length zs" shows "bitenc zs ! (2 * i) = \" and "bitenc zs ! (2 * i + 1) = zs ! i" proof - let ?f = "\h. [\, h]" let ?xs = "concat (map ?f zs)" have eqtake: "bitenc (take i zs) = take (2 * i) (bitenc zs)" if "i \ length zs" for i zs proof - have "take (2 * i) (bitenc zs) = take (2 * i) (bitenc (take i zs @ drop i zs))" by simp then have "take (2 * i) (bitenc zs) = take (2 * i) (bitenc (take i zs) @ (bitenc (drop i zs)))" by (metis concat_append map_append) then show ?thesis using length_bitenc that by simp qed have eqdrop: "bitenc (drop i zs) = drop (2 * i) (bitenc zs)" if "i < length zs" for i proof - have "drop (2 * i) (bitenc zs) = drop (2 * i) (bitenc (take i zs @ drop i zs))" by simp then have "drop (2 * i) (bitenc zs) = drop (2 * i) (bitenc (take i zs) @ bitenc (drop i zs))" by (metis concat_append map_append) then show ?thesis using length_bitenc that by simp qed have take2: "take 2 (drop (2 * i) (bitenc zs)) = ?f (zs ! i)" if "i < length zs" for i proof - have 1: "1 \ length (drop i zs)" using that by simp have "take 2 (drop (2*i) (bitenc zs)) = take 2 (bitenc (drop i zs))" using that eqdrop by simp also have "... = bitenc (take 1 (drop i zs))" using 1 eqtake by simp also have "... = bitenc [zs ! i]" using that by (metis Cons_nth_drop_Suc One_nat_def take0 take_Suc_Cons) also have "... = ?f (zs ! i)" by simp finally show ?thesis . qed show "bitenc zs ! (2 * i) = \" proof - have "bitenc zs ! (2 * i) = drop (2 * i) (bitenc zs) ! 0" using assms drop0 length_bitenc by simp also have "... = take 2 (drop (2 * i) (bitenc zs)) ! 0" using eqdrop by simp also have "... = ?f (zs ! i) ! 0" using assms take2 by simp also have "... = \" by simp finally show ?thesis . qed show "bitenc zs ! (2*i + 1) = zs ! i" proof - have "bitenc zs ! (2*i+1) = drop (2 * i) (bitenc zs) ! 1" using assms length_bitenc by simp also have "... = take 2 (drop (2*i) (bitenc zs)) ! 1" using eqdrop by simp also have "... = ?f (zs ! i) ! 1" using assms(1) take2 by simp also have "... = zs ! i" by simp finally show ?thesis . qed qed lemma string_pair_first_nth: assumes "i < length x" shows "\x, y\ ! (2 * i) = \" and "\x, y\ ! (2 * i + 1) = x ! i" proof - have "\x, y\ ! (2*i) = concat (map (\h. [\, h]) x) ! (2*i)" using string_pair_def length_bitenc by (simp add: assms nth_append) then show "\x, y\ ! (2 * i) = \" using bitenc_nth(1) assms by simp have "2 * i + 1 < 2 * length x" using assms by simp then have "\x, y\ ! (2*i+1) = concat (map (\h. [\, h]) x) ! (2*i+1)" using string_pair_def length_bitenc[of x] assms nth_append by force then show "\x, y\ ! (2 * i + 1) = x ! i" using bitenc_nth(2) assms by simp qed lemma string_pair_sep_nth: shows "\x, y\ ! (2 * length x) = \" and "\x, y\ ! (2 * length x + 1) = \" using string_pair_def length_bitenc by (metis append_Cons nth_append_length) (simp add: length_bitenc nth_append string_pair_def) lemma string_pair_second_nth: assumes "i < length y" shows "\x, y\ ! (2 * length x + 2 + 2 * i) = \" and "\x, y\ ! (2 * length x + 2 + 2 * i + 1) = y ! i" proof - have "\x, y\ ! (2 * length x + 2 + 2*i) = concat (map (\h. [\, h]) y) ! (2*i)" using string_pair_def length_bitenc by (simp add: assms nth_append) then show "\x, y\ ! (2 * length x + 2 + 2 * i) = \" using bitenc_nth(1) assms by simp have "2 * i + 1 < 2 * length y" using assms by simp then have "\x, y\ ! (2 * length x + 2 + 2*i+1) = concat (map (\h. [\, h]) y) ! (2*i+1)" using string_pair_def length_bitenc[of x] assms nth_append by force then show "\x, y\ ! (2 * length x + 2 + 2 * i + 1) = y ! i" using bitenc_nth(2) assms by simp qed lemma string_pair_inj: assumes "\x1, y1\ = \x2, y2\" shows "x1 = x2 \ y1 = y2" proof show "x1 = x2" proof (rule ccontr) assume neq: "x1 \ x2" consider "length x1 = length x2" | "length x1 < length x2" | "length x1 > length x2" by linarith then show False proof (cases) case 1 then obtain i where i: "i < length x1" "x1 ! i \ x2 ! i" using neq list_eq_iff_nth_eq by blast then have "\x1, y1\ ! (2 * i + 1) = x1 ! i" and "\x2, y2\ ! (2 * i + 1) = x2 ! i" using 1 string_pair_first_nth by simp_all then show False using assms i(2) by simp next case 2 let ?i = "length x1" have "\x1, y1\ ! (2 * ?i) = \" using string_pair_sep_nth by simp moreover have "\x2, y2\ ! (2 * ?i) = \" using string_pair_first_nth 2 by simp ultimately show False using assms by simp next case 3 let ?i = "length x2" have "\x2, y2\ ! (2 * ?i) = \" using string_pair_sep_nth by simp moreover have "\x1, y1\ ! (2 * ?i) = \" using string_pair_first_nth 3 by simp ultimately show False using assms by simp qed qed then have len_x_eq: "length x1 = length x2" by simp then have len_y_eq: "length y1 = length y2" using assms length_string_pair by (smt (verit) Suc_1 Suc_mult_cancel1 add_left_imp_eq add_right_cancel) show "y1 = y2" proof (rule ccontr) assume neq: "y1 \ y2" then obtain i where i: "i < length y1" "y1 ! i \ y2 ! i" using list_eq_iff_nth_eq len_y_eq by blast then have "\x1, y1\ ! (2 * length x1 + 2 + 2 * i + 1) = y1 ! i" and "\x2, y2\ ! (2 * length x2 + 2 + 2 * i + 1) = y2 ! i" using string_pair_second_nth len_y_eq by simp_all then show False using assms i(2) len_x_eq by simp qed qed text \ Turing machines have to deal with pairs of symbol sequences rather than strings. \ abbreviation pair :: "string \ string \ symbol list" ("\_; _\") where "\x; y\ \ string_to_symbols \x, y\" lemma symbols_lt_pair: "symbols_lt 4 \x; y\" by simp lemma length_pair: "length \x; y\ = 2 * length x + 2 * length y + 2" by (simp add: length_string_pair) lemma pair_inj: assumes "\x1; y1\ = \x2; y2\" shows "x1 = x2 \ y1 = y2" using string_pair_inj assms symbols_to_string_to_symbols by metis subsection \Big-Oh and polynomials\label{s:tm-basic-bigoh}\ text \ The Big-Oh notation is standard~\cite[Definition~0.2]{ccama}. It can be defined with $c$ ranging over real or natural numbers. We choose natural numbers for simplicity. \ definition big_oh :: "(nat \ nat) \ (nat \ nat) \ bool" where "big_oh g f \ \c m. \n>m. g n \ c * f n" text \ Some examples: \ proposition "big_oh (\n. n) (\n. n)" using big_oh_def by auto proposition "big_oh (\n. n) (\n. n * n)" using big_oh_def by auto proposition "big_oh (\n. 42 * n) (\n. n * n)" proof- have "\n>0::nat. 42 * n \ 42 * n * n" by simp then have "\(c::nat)>0. \n>0. 42 * n \ c * n * n" using zero_less_numeral by blast then show ?thesis using big_oh_def by auto qed proposition "\ big_oh (\n. n * n) (\n. n)" (is "\ big_oh ?g ?f") proof assume "big_oh (\n. n * n) (\n. n)" then obtain c m where "\n>m. ?g n \ c * ?f n" using big_oh_def by auto then have 1: "\n>m. n * n \ c * n" by auto define nn where "nn = max (m + 1) (c + 1)" then have 2: "nn > m" by simp then have "nn * nn > c * nn" by (simp add: nn_def max_def) with 1 2 show False using not_le by blast qed text \ Some lemmas helping with polynomial upper bounds. \ lemma pow_mono: fixes n d1 d2 :: nat assumes "d1 \ d2" and "n > 0" shows "n ^ d1 \ n ^ d2" using assms by (simp add: Suc_leI power_increasing) lemma pow_mono': fixes n d1 d2 :: nat assumes "d1 \ d2" and "0 < d1" shows "n ^ d1 \ n ^ d2" using assms by (metis dual_order.eq_iff less_le_trans neq0_conv pow_mono power_eq_0_iff) lemma linear_le_pow: fixes n d1 :: nat assumes "0 < d1" shows "n \ n ^ d1" using assms by (metis One_nat_def gr_implies_not0 le_less_linear less_Suc0 self_le_power) text \ The next definition formalizes the phrase ``polynomially bounded'' and the term ``polynomial'' in ``polynomial running-time''. This is often written ``$f(n) = n^{O(1)}$'' (for example, Arora and Barak~\cite[Example 0.3]{ccama}). \ definition big_oh_poly :: "(nat \ nat) \ bool" where "big_oh_poly f \ \d. big_oh f (\n. n ^ d)" lemma big_oh_poly: "big_oh_poly f \ (\d c n\<^sub>0. \n>n\<^sub>0. f n \ c * n ^ d)" using big_oh_def big_oh_poly_def by auto lemma big_oh_polyI: assumes "\n. n > n\<^sub>0 \ f n \ c * n ^ d" shows "big_oh_poly f" using assms big_oh_poly by auto lemma big_oh_poly_const: "big_oh_poly (\n. c)" proof - let ?c = "max 1 c" have "(\n. c) n \ ?c * n ^ 1" if "n > 0" for n proof - have "c \ n * ?c" by (metis (no_types) le_square max.cobounded2 mult.assoc mult_le_mono nat_mult_le_cancel_disj that) then show ?thesis by (simp add: mult.commute) qed then show ?thesis using big_oh_polyI[of 0 _ ?c] by simp qed lemma big_oh_poly_poly: "big_oh_poly (\n. n ^ d)" using big_oh_polyI[of 0 _ 1 d] by simp lemma big_oh_poly_id: "big_oh_poly (\n. n)" using big_oh_poly_poly[of 1] by simp lemma big_oh_poly_le: assumes "big_oh_poly f" and "\n. g n \ f n" shows "big_oh_poly g" using assms big_oh_polyI by (metis big_oh_poly le_trans) lemma big_oh_poly_sum: assumes "big_oh_poly f1" and "big_oh_poly f2" shows "big_oh_poly (\n. f1 n + f2 n)" proof- obtain d1 c1 m1 where 1: "\n>m1. f1 n \ c1 * n ^ d1" using big_oh_poly assms(1) by blast obtain d2 c2 m2 where 2: "\n>m2. f2 n \ c2 * n ^ d2" using big_oh_poly assms(2) by blast let ?f3 = "\n. f1 n + f2 n" let ?c3 = "max c1 c2" let ?m3 = "max m1 m2" let ?d3 = "max d1 d2" have "\n>?m3. f1 n \ ?c3 * n ^ d1" using 1 by (simp add: max.coboundedI1 nat_mult_max_left) moreover have "\n>?m3. n ^ d1 \ n^?d3" using pow_mono by simp ultimately have *: "\n>?m3. f1 n \ ?c3 * n^?d3" using order_subst1 by fastforce have "\n>?m3. f2 n \ ?c3 * n ^ d2" using 2 by (simp add: max.coboundedI2 nat_mult_max_left) moreover have "\n>?m3. n ^ d2 \ n ^ ?d3" using pow_mono by simp ultimately have "\n>?m3. f2 n \ ?c3 * n ^ ?d3" using order_subst1 by fastforce then have "\n>?m3. f1 n + f2 n \ ?c3 * n ^ ?d3 + ?c3 * n ^ ?d3" using * by fastforce then have "\n>?m3. f1 n + f2 n \ 2 * ?c3 * n ^ ?d3" by auto then have "\d c m. \n>m. ?f3 n \ c * n ^ d" by blast then show ?thesis using big_oh_poly by simp qed lemma big_oh_poly_prod: assumes "big_oh_poly f1" and "big_oh_poly f2" shows "big_oh_poly (\n. f1 n * f2 n)" proof- obtain d1 c1 m1 where 1: "\n>m1. f1 n \ c1 * n ^ d1" using big_oh_poly assms(1) by blast obtain d2 c2 m2 where 2: "\n>m2. f2 n \ c2 * n ^ d2" using big_oh_poly assms(2) by blast let ?f3 = "\n. f1 n * f2 n" let ?c3 = "max c1 c2" let ?m3 = "max m1 m2" have "\n>?m3. f1 n \ ?c3 * n ^ d1" using 1 by (simp add: max.coboundedI1 nat_mult_max_left) moreover have "\n>?m3. n ^ d1 \ n ^ d1" using pow_mono by simp ultimately have *: "\n>?m3. f1 n \ ?c3 * n ^ d1" using order_subst1 by fastforce have "\n>?m3. f2 n \ ?c3 * n ^ d2" using 2 by (simp add: max.coboundedI2 nat_mult_max_left) moreover have "\n>?m3. n ^ d2 \ n ^ d2" using pow_mono by simp ultimately have "\n>?m3. f2 n \ ?c3 * n ^ d2" using order_subst1 by fastforce then have "\n>?m3. f1 n * f2 n \ ?c3 * n ^ d1 * ?c3 * n ^ d2" using * mult_le_mono by (metis mult.assoc) then have "\n>?m3. f1 n * f2 n \ ?c3 * ?c3 * n ^ d1 * n ^ d2" by (simp add: semiring_normalization_rules(16)) then have "\n>?m3. f1 n * f2 n \ ?c3 * ?c3 * n ^ (d1 + d2)" by (simp add: mult.assoc power_add) then have "\d c m. \n>m. ?f3 n \ c * n ^ d" by blast then show ?thesis using big_oh_poly by simp qed lemma big_oh_poly_offset: assumes "big_oh_poly f" shows "\b c d. d > 0 \ (\n. f n \ b + c * n ^ d)" proof - obtain d c m where dcm: "\n>m. f n \ c * n ^ d" using assms big_oh_poly by auto have *: "f n \ c * n ^ Suc d" if "n > m" for n proof - have "n > 0" using that by simp then have "n ^ d \ n ^ Suc d" by simp then have "c * n ^ d \ c * n ^ Suc d" by simp then show "f n \ c * n ^ Suc d" using dcm order_trans that by blast qed define b :: nat where "b = Max {f n | n. n \ m}" then have "y \ b" if "y \ {f n | n. n \ m}" for y using that by simp then have "f n \ b" if "n \ m" for n using that by auto then have "f n \ b + c * n ^ Suc d" for n using * by (meson trans_le_add1 trans_le_add2 verit_comp_simplify1(3)) then show ?thesis using * dcm(1) by blast qed lemma big_oh_poly_composition: assumes "big_oh_poly f1" and "big_oh_poly f2" shows "big_oh_poly (f2 \ f1)" proof- obtain d1 c1 m1 where 1: "\n>m1. f1 n \ c1 * n ^ d1" using big_oh_poly assms(1) by blast obtain d2 c2 b where 2: "\n. f2 n \ b + c2 * n ^ d2" using big_oh_poly_offset assms(2) by blast define c where "c = c2 * c1 ^ d2" have 3: "\n>m1. f1 n \ c1 * n ^ d1" using 1 by simp have "\n>m1. f2 n \ b + c2 * n ^ d2" using 2 by simp { fix n assume "n > m1" then have 4: "(f1 n) ^ d2 \ (c1 * n ^ d1) ^ d2" using 3 by (simp add: power_mono) have "f2 (f1 n) \ b + c2 * (f1 n) ^ d2" using 2 by simp also have "... \ b + c2 * (c1 * n ^ d1) ^ d2" using 4 by simp also have "... = b + c2 * c1 ^ d2 * n ^ (d1 * d2)" by (simp add: power_mult power_mult_distrib) also have "... = b + c * n ^ (d1 * d2)" using c_def by simp also have "... \ b * n ^ (d1 * d2) + c * n ^ (d1 * d2)" using `n > m1` by simp also have "... \ (b + c) * n ^ (d1 * d2)" by (simp add: comm_semiring_class.distrib) finally have "f2 (f1 n) \ (b + c) * n ^ (d1 * d2)" . } then show ?thesis using big_oh_polyI[of m1 _ "b + c" "d1 * d2"] by simp qed lemma big_oh_poly_pow: fixes f :: "nat \ nat" and d :: nat assumes "big_oh_poly f" shows "big_oh_poly (\n. f n ^ d)" proof - let ?g = "\n. n ^ d" have "big_oh_poly ?g" using big_oh_poly_poly by simp moreover have "(\n. f n ^ d) = ?g \ f" by auto ultimately show ?thesis using assms big_oh_poly_composition by simp qed text \ The textbook does not give an explicit definition of polynomials. It treats them as functions between natural numbers. So assuming the coefficients are natural numbers too, seems natural. We justify this choice when defining $\NP$ in Section~\ref{s:TC-NP}. \null \ definition polynomial :: "(nat \ nat) \ bool" where "polynomial f \ \cs. \n. f n = (\i\[0.._. c)" proof - let ?cs = "[c]" have "\n. (\_. c) n = (\i\[0..n::nat. id n = (\i\[0.. nat" assumes "polynomial f" shows "big_oh_poly f" proof - have "big_oh_poly (\n. (\i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..i\[0..n. \i\[0..n. (\i\[0..n. cs ! len * n ^ len)" using big_oh_poly_poly big_oh_poly_prod big_oh_poly_const by simp moreover have "big_oh_poly (\n. (\i\[0..n. \i\[0..n. (\i\[0..Increasing the alphabet or the number of tapes\label{s:tm-trans}\ text \ For technical reasons it is sometimes necessary to add tapes to a machine or to formally enlarge its alphabet such that it matches another machine's tape number or alphabet size without changing the behavior of the machine. The primary use of this is when composing machines with unequal alphabets or tape numbers (see Section~\ref{s:tm-composing}). \ subsection \Enlarging the alphabet\ text \ A Turing machine over alphabet $G$ is not necessarily a Turing machine over a larger alphabet $G' > G$ because reading a symbol in $\{G, \dots, G'-1\}$ the TM may write a symbol $\geq G'$. This is easy to remedy by modifying the TM to do nothing when it reads a symbol $\geq G$. It then formally satisfies the alphabet restriction property of Turing commands. This is rather crude, because the new TM loops infinitely on encountering a ``forbidden'' symbol, but it is good enough for our purposes. The next function performs this transformation on a TM $M$ over alphabet $G$. The resulting machine is a Turing machine for every alphabet size $G' \ge G$. \ definition enlarged :: "nat \ machine \ machine" where "enlarged G M \ map (\cmd rs. if symbols_lt G rs then cmd rs else (0, map (\r. (r, Stay)) rs)) M" lemma length_enlarged: "length (enlarged G M) = length M" using enlarged_def by simp lemma enlarged_nth: assumes "symbols_lt G gs" and "i < length M" shows "(M ! i) gs = (enlarged G M ! i) gs" using assms enlarged_def by simp lemma enlarged_write: assumes "length gs = k" and "i < length M" and "turing_machine k G M" shows "length (snd ((M ! i) gs)) = length (snd ((enlarged G M ! i) gs))" proof (cases "symbols_lt G gs") case True then show ?thesis using assms enlarged_def by simp next case False then have "(enlarged G M ! i) gs = (0, map (\r. (r, Stay)) gs)" using assms enlarged_def by auto then show ?thesis using assms turing_commandD(1) turing_machine_def by (metis length_map nth_mem snd_conv) qed lemma turing_machine_enlarged: assumes "turing_machine k G M" and "G' \ G" shows "turing_machine k G' (enlarged G M)" proof let ?M = "enlarged G M" show "2 \ k" and "4 \ G'" using assms turing_machine_def by simp_all show "turing_command k (length ?M) G' (?M ! i)" if i: "i < length ?M" for i proof have len: "length ?M = length M" using enlarged_def by simp then have 1: "turing_command k (length M) G (M ! i)" using assms(1) that turing_machine_def by simp show "\gs. length gs = k \ length ([!!] (?M ! i) gs) = length gs" using enlarged_write that 1 len assms(1) by (metis turing_commandD(1)) show "(?M ! i) gs [.] j < G'" if "length gs = k" "(\i. i < length gs \ gs ! i < G')" "j < length gs" for gs j proof (cases "symbols_lt G gs") case True then have "(?M ! i) gs = (M ! i) gs" using enlarged_def i by simp moreover have "(M ! i) gs [.] j < G" using "1" turing_commandD(2) that(1,3) True by simp ultimately show ?thesis using assms(2) by simp next case False then have "(?M ! i) gs = (0, map (\r. (r, Stay)) gs)" using enlarged_def i by auto then show ?thesis using that by simp qed show "(?M ! i) gs [.] 0 = gs ! 0" if "length gs = k" and "k > 0" for gs proof (cases "symbols_lt G gs") case True then show ?thesis using enlarged_def i "1" turing_command_def that by simp next case False then have "(?M ! i) gs = (0, map (\r. (r, Stay)) gs)" using that enlarged_def i by auto then show ?thesis using assms(1) turing_machine_def that by simp qed show "[*] ((?M ! i) gs) \ length ?M" if "length gs = k" for gs proof (cases "symbols_lt G gs") case True then show ?thesis using enlarged_def i that assms(1) turing_machine_def "1" turing_commandD(4) enlarged_nth len by (metis (no_types, lifting)) next case False then show ?thesis using that enlarged_def i by auto qed qed qed text \ The enlarged machine has the same behavior as the original machine when started on symbols over the original alphabet $G$. \ lemma execute_enlarged: assumes "turing_machine k G M" and "symbols_lt G zs" shows "execute (enlarged G M) (start_config k zs) t = execute M (start_config k zs) t" proof (induction t) case 0 then show ?case by simp next case (Suc t) let ?M = "enlarged G M" have "execute ?M (start_config k zs) (Suc t) = exe ?M (execute ?M (start_config k zs) t)" by simp also have "... = exe ?M (execute M (start_config k zs) t)" (is "_ = exe ?M ?cfg") using Suc by simp also have "... = execute M (start_config k zs) (Suc t)" proof (cases "fst ?cfg < length M") case True then have "exe ?M ?cfg = sem (?M ! (fst ?cfg)) ?cfg" (is "_ = sem ?cmd ?cfg") using exe_lt_length length_enlarged by simp then have "exe ?M ?cfg = (fst (?cmd (config_read ?cfg)), map (\(a, tp). act a tp) (zip (snd (?cmd (config_read ?cfg))) (snd ?cfg)))" using sem' by simp moreover have "symbols_lt G (config_read ?cfg)" using read_alphabet' assms by auto ultimately have "exe ?M ?cfg = (fst ((M ! (fst ?cfg)) (config_read ?cfg)), map (\(a, tp). act a tp) (zip (snd ((M ! (fst ?cfg)) (config_read ?cfg))) (snd ?cfg)))" using True enlarged_nth by auto then have "exe ?M ?cfg = exe M ?cfg" using sem' by (simp add: True exe_lt_length) then show ?thesis using Suc by simp next case False then show ?thesis using Suc enlarged_def exe_def by auto qed finally show ?case . qed lemma transforms_enlarged: assumes "turing_machine k G M" and "symbols_lt G zs" and "transforms M (snd (start_config k zs)) t tps1" shows "transforms (enlarged G M) (snd (start_config k zs)) t tps1" proof - let ?tps = "snd (start_config k zs)" have "\t'\t. execute M (start_config k zs) t' = (length M, tps1)" using assms(3) transforms_def transits_def start_config_def by simp then have "\t'\t. execute (enlarged G M) (start_config k zs) t' = (length M, tps1)" using assms(1,2) transforms_def transits_def execute_enlarged by auto moreover have "length M = length (enlarged G M)" using enlarged_def by simp ultimately show ?thesis using start_config_def transforms_def transitsI by auto qed subsection \Increasing the number of tapes\ text \ We can add tapes to a Turing machine in such a way that on the additional tapes the machine does nothing. While the new tapes could go anywhere, we only consider appending them at the end or inserting them at the beginning. \ subsubsection \Appending tapes at the end\ text \ The next function turns a $k$-tape Turing machine into a $k'$-tape Turing machine (for $k' \geq k$) by appending $k' - k$ tapes at the end. \ definition append_tapes :: "nat \ nat \ machine \ machine" where "append_tapes k k' M \ map (\cmd rs. (fst (cmd (take k rs)), snd (cmd (take k rs)) @ (map (\i. (rs ! i, Stay)) [k..j. (gs ! j, Stay)) [k.. k" shows "turing_machine k' G (append_tapes k k' M)" proof let ?M = "append_tapes k k' M" show "2 \ k'" using assms turing_machine_def by simp show "4 \ G" using assms(1) turing_machine_def by simp show "turing_command k' (length ?M) G (?M ! i)" if "i < length ?M" for i proof have "i < length M" using that by (simp add: append_tapes_def) then have turing_command: "turing_command k (length M) G (M ! i)" using assms(1) that turing_machine_def by simp have ith: "append_tapes k k' M ! i = (\rs. (fst ((M ! i) (take k rs)), snd ((M ! i) (take k rs)) @ (map (\j. (rs ! j, Stay)) [k..gs. length gs = k' \ length ([!!] (append_tapes k k' M ! i) gs) = length gs" using assms(2) ith turing_command turing_commandD by simp show "(append_tapes k k' M ! i) gs [.] j < G" if "length gs = k'" "\i. i < length gs \ gs ! i < G" "j < length gs" for j gs proof (cases "j < k") case True let ?gs = "take k gs" have len: "length ?gs = k" using that(1) assms(2) by simp have "\i. i < length ?gs \ ?gs ! i < G" using that(2) by simp then have "\i' k" by simp have *: "length (snd ((M ! i) (take k gs))) = k" using turing_commandD(1)[OF turing_command] assms(2) that(1) by auto have "(append_tapes k k' M ! i) gs [.] j = fst ((snd ((M ! i) (take k gs)) @ (map (\j. (gs ! j, Stay)) [k..j. (gs ! j, Stay)) [k.. k` by (simp add: False nth_append) also have "... = fst (gs ! j, Stay)" by (metis False \k \ j\ add_diff_inverse_nat diff_less_mono length_upt nth_map nth_upt that(1,3)) also have "... = gs ! j" by simp also have "... < G" using that(2,3) by simp finally show ?thesis by simp qed show "(append_tapes k k' M ! i) gs [.] 0 = gs ! 0" if "length gs = k'" for gs proof - have "k > 0" using assms(1) turing_machine_def by simp then have 1: "(M ! i) rs [.] 0 = rs ! 0" if "length rs = k" for rs using turing_commandD(3)[OF turing_command that] that by simp have len: "length (take k gs) = k" by (simp add: assms(2) min_absorb2 that(1)) then have *: "length (snd ((M ! i) (take k gs))) = k" using turing_commandD(1)[OF turing_command] by auto have "(append_tapes k k' M ! i) gs [.] 0 = fst ((snd ((M ! i) (take k gs)) @ (map (\j. (gs ! j, Stay)) [k..0 < k\ by simp qed show "[*] ((append_tapes k k' M ! i) gs) \ length (append_tapes k k' M)" if "length gs = k'" for gs proof - have "length (take k gs) = k" using assms(2) that by simp then have 1: "fst ((M ! i) (take k gs)) \ length M" using turing_commandD[OF turing_command] \i < length M\ assms(1) turing_machine_def by blast moreover have "fst ((append_tapes k k' M ! i) gs) = fst ((M ! i) (take k gs))" using ith by simp ultimately show "fst ((append_tapes k k' M ! i) gs) \ length (append_tapes k k' M)" using length_append_tapes by metis qed qed qed lemma execute_append_tapes: assumes "turing_machine k G M" and "k' \ k" and "length tps = k'" shows "execute (append_tapes k k' M) (q, tps) t = (fst (execute M (q, take k tps) t), snd (execute M (q, take k tps) t) @ drop k tps)" proof (induction t) case 0 then show ?case by simp next case (Suc t) let ?M = "append_tapes k k' M" let ?cfg = "execute M (q, take k tps) t" let ?cfg' = "execute M (q, take k tps) (Suc t)" have "execute ?M (q, tps) (Suc t) = exe ?M (execute ?M (q, tps) t)" by simp also have "... = exe ?M (fst ?cfg, snd ?cfg @ drop k tps)" using Suc by simp also have "... = (fst ?cfg', snd ?cfg' @ drop k tps)" proof (cases "fst ?cfg < length ?M") case True have "sem (?M ! (fst ?cfg)) (fst ?cfg, snd ?cfg @ drop k tps) = (fst ?cfg', snd ?cfg' @ drop k tps)" proof (rule semI) have "turing_machine k' G (append_tapes k k' M)" using append_tapes_tm[OF assms(1,2)] by simp then show 1: "proper_command k' (append_tapes k k' M ! fst (execute M (q, take k tps) t))" using True turing_machine_def turing_commandD by (metis nth_mem) show 2: "length (snd ?cfg @ drop k tps) = k'" using assms execute_num_tapes by fastforce show "length (snd ?cfg' @ drop k tps) = k'" by (metis (no_types, lifting) append_take_drop_id assms execute_num_tapes length_append length_take min_absorb2 snd_conv) show "fst ((?M ! fst ?cfg) (read (snd ?cfg @ drop k tps))) = fst ?cfg'" proof - have less': "fst ?cfg < length M" using True by (simp add: length_append_tapes) let ?tps = "snd ?cfg @ drop k tps" have "length (snd ?cfg) = k" using assms execute_num_tapes by fastforce then have take2: "take k ?tps = snd ?cfg" by simp let ?rs = "read ?tps" have len: "length ?rs = k'" using 2 read_length by simp have take2': "take k ?rs = read (snd ?cfg)" using read_def take2 by (metis (mono_tags, lifting) take_map) have "fst ((?M ! fst ?cfg) ?rs) = fst (fst ((M ! fst ?cfg) (take k ?rs)), snd ((M ! fst ?cfg) (take k ?rs)) @ (map (\j. (?rs ! j, Stay)) [k..j. (?rs ! j, Stay)) [k..j. (?rs ! j, Stay)) [k..j. (?rs ! j, Stay)) [k.. k" by simp have len': "length (snd ?cfg') = k" using assms(1) exe_num_tapes len2 by auto have rs: "?rs ! j = read tps ! j" using tps by (metis (no_types, lifting) "2" assms(3) that nth_map read_def) have "act ((snd ((M ! fst ?cfg) (read (snd ?cfg))) @ (map (\j. (?rs ! j, Stay)) [k..j. (?rs ! j, Stay)) [k..Inserting tapes at the beginning\ text \ The next function turns a $k$-tape Turing machine into a $(k + d)$-tape Turing machine by inserting $d$ tapes at the beginning. \ definition prepend_tapes :: "nat \ machine \ machine" where "prepend_tapes d M \ map (\cmd rs. (fst (cmd (drop d rs)), map (\h. (h, Stay)) (take d rs) @ snd (cmd (drop d rs)))) M" lemma prepend_tapes_at: assumes "i < length M" shows "(prepend_tapes d M ! i) gs = (fst ((M ! i) (drop d gs)), map (\h. (h, Stay)) (take d gs) @ snd ((M ! i) (drop d gs)))" using assms prepend_tapes_def by simp lemma prepend_tapes_tm: assumes "turing_machine k G M" shows "turing_machine (d + k) G (prepend_tapes d M)" proof show "2 \ d + k" using assms turing_machine_def by simp show "4 \ G" using assms turing_machine_def by simp let ?M = "prepend_tapes d M" show "turing_command (d + k) (length ?M) G (?M ! i)" if "i < length ?M" for i proof have len: "i < length M" using that prepend_tapes_def by simp then have *: "(?M ! i) gs = (fst ((M ! i) (drop d gs)), map (\h. (h, Stay)) (take d gs) @ snd ((M ! i) (drop d gs)))" if "length gs = d + k" for gs using prepend_tapes_def that by simp have tc: "turing_command k (length M) G (M ! i)" using that turing_machine_def len assms by simp show "length (snd ((?M ! i) gs)) = length gs" if "length gs = d + k" for gs using * that turing_commandD[OF tc] by simp show "(?M ! i) gs [.] j < G" if "length gs = d + k" "(\i. i < length gs \ gs ! i < G)" "j < length gs" for gs j proof (cases "j < d") case True have "(?M ! i) gs [.] j = fst ((map (\h. (h, Stay)) (take d gs) @ snd ((M ! i) (drop d gs))) ! j)" using * that(1) by simp also have "... = fst (map (\h. (h, Stay)) (take d gs) ! j)" using True that(1) by (simp add: nth_append) also have "... = gs ! j" by (simp add: True that(3)) finally have "(?M ! i) gs [.] j = gs ! j" . then show ?thesis using that(2,3) by simp next case False have "(?M ! i) gs [.] j = fst ((map (\h. (h, Stay)) (take d gs) @ snd ((M ! i) (drop d gs))) ! j)" using * that(1) by simp also have "... = fst (snd ((M ! i) (drop d gs)) ! (j - d))" using False that(1) by (metis (no_types, lifting) add_diff_cancel_left' append_take_drop_id diff_add_inverse2 length_append length_drop length_map nth_append) also have "... < G" using False that turing_commandD[OF tc] by simp finally show ?thesis by simp qed show "(?M ! i) gs [.] 0 = gs ! 0" if "length gs = d + k" and "d + k > 0" for gs proof (cases "d = 0") case True then have "(?M ! i) gs [.] 0 = fst (snd ((M ! i) gs) ! 0)" using * that(1) by simp then show ?thesis using True that turing_commandD[OF tc] by simp next case False then have "(?M ! i) gs [.] 0 = fst ((map (\h. (h, Stay)) (take d gs)) ! 0)" using * that(1) by (simp add: nth_append) also have "... = fst ((map (\h. (h, Stay)) gs) ! 0)" using False by (metis gr_zeroI nth_take take_map) also have "... = gs ! 0" using False that by simp finally show ?thesis by simp qed show "[*] ((?M ! i) gs) \ length ?M" if "length gs = d + k" for gs proof - have "fst ((?M ! i) gs) = fst ((M ! i) (drop d gs))" using that * by simp moreover have "length (drop d gs) = k" using that by simp ultimately have "fst ((?M ! i) gs) \ length M" using turing_commandD(4)[OF tc] by fastforce then show "fst ((?M ! i) gs) \ length ?M" using prepend_tapes_def by simp qed qed qed definition shift_cfg :: "tape list \ config \ config" where "shift_cfg tps cfg \ (fst cfg, tps @ snd cfg)" lemma execute_prepend_tapes: assumes "turing_machine k G M" and "length tps = d" and "||cfg0|| = k" shows "execute (prepend_tapes d M) (shift_cfg tps cfg0) t = shift_cfg tps (execute M cfg0 t)" proof (induction t) case 0 show ?case by simp next case (Suc t) let ?M = "prepend_tapes d M" let ?scfg = "shift_cfg tps cfg0" let ?scfg' = "execute ?M ?scfg t" let ?cfg' = "execute M cfg0 t" have fst: "fst ?cfg' = fst ?scfg'" using shift_cfg_def Suc.IH by simp have len: "||?cfg'|| = k" using assms(1,3) execute_num_tapes read_length by auto have len_s: "||?scfg'|| = d + k" using prepend_tapes_tm[OF assms(1)] shift_cfg_def assms(2,3) execute_num_tapes read_length by (metis length_append snd_conv) let ?srs = "read (snd ?scfg')" let ?rs = "read (snd ?cfg')" have len_rs: "length ?rs = k" using assms(1,3) execute_num_tapes read_length by auto moreover have len_srs: "length ?srs = k + d" using prepend_tapes_tm[OF assms(1)] shift_cfg_def assms(2,3) by (metis add.commute execute_num_tapes length_append read_length snd_conv) ultimately have srs_rs: "drop d ?srs = ?rs" using Suc shift_cfg_def read_def by simp have *: "execute ?M ?scfg (Suc t) = exe ?M ?scfg'" by simp show ?case proof (cases "fst ?scfg' \ length ?M") case True then show ?thesis using * Suc exe_ge_length shift_cfg_def prepend_tapes_def by auto next case running: False then have scmd: "?M ! (fst ?scfg') = (\gs. (fst ((M ! (fst ?scfg')) (drop d gs)), map (\h. (h, Stay)) (take d gs) @ snd ((M ! (fst ?scfg')) (drop d gs))))" (is "?scmd = _") using prepend_tapes_at prepend_tapes_def by auto then have cmd: "?M ! (fst ?scfg') = (\gs. (fst ((M ! (fst ?cfg')) (drop d gs)), map (\h. (h, Stay)) (take d gs) @ snd ((M ! (fst ?cfg')) (drop d gs))))" using fst by simp let ?cmd = "M ! (fst ?cfg')" have "execute ?M ?scfg (Suc t) = sem (?M ! (fst ?scfg')) ?scfg'" using running * exe_lt_length by simp then have lhs: "execute ?M ?scfg (Suc t) = (fst (?scmd ?srs), map (\(a, tp). act a tp) (zip (snd (?scmd ?srs)) (snd ?scfg')))" (is "_ = ?lhs") using sem' by simp have "shift_cfg tps (execute M cfg0 (Suc t)) = shift_cfg tps (exe M ?cfg')" by simp also have "... = shift_cfg tps (sem (M ! (fst ?cfg')) ?cfg')" using exe_lt_length running fst prepend_tapes_def by auto also have "... = shift_cfg tps (fst (?cmd ?rs), map (\(a, tp). act a tp) (zip (snd (?cmd ?rs)) (snd ?cfg')))" using sem' by simp also have "... = (fst (?cmd ?rs), tps @ map (\(a, tp). act a tp) (zip (snd (?cmd ?rs)) (snd ?cfg')))" using shift_cfg_def by simp finally have rhs: "shift_cfg tps (execute M cfg0 (Suc t)) = (fst (?cmd ?rs), tps @ map (\(a, tp). act a tp) (zip (snd (?cmd ?rs)) (snd ?cfg')))" (is "_ = ?rhs") . have "?lhs = ?rhs" proof standard+ show "fst (?scmd ?srs) = fst (?cmd ?rs)" using srs_rs cmd by simp show "map (\(a, tp). act a tp) (zip (snd (?scmd ?srs)) (snd ?scfg')) = tps @ map (\(a, tp). act a tp) (zip (snd (?cmd ?rs)) (snd ?cfg'))" (is "?l = ?r") proof (rule nth_equalityI) have lenl: "length ?l = d + k" using lhs execute_num_tapes assms prepend_tapes_tm len_s - by (smt (z3) length_append shift_cfg_def snd_conv) + by (smt (verit) length_append shift_cfg_def snd_conv) moreover have "length ?r = d + k" using rhs execute_num_tapes assms shift_cfg_def by (metis (mono_tags, lifting) length_append snd_conv) ultimately show "length ?l = length ?r" by simp show "?l ! j = ?r ! j" if "j < length ?l" for j proof (cases "j < d") case True let ?at = "zip (snd (?scmd ?srs)) (snd ?scfg') ! j" have "?l ! j = act (fst ?at) (snd ?at)" using that by simp moreover have "fst ?at = snd (?scmd ?srs) ! j" using that by simp moreover have "snd ?at = snd ?scfg' ! j" using that by simp ultimately have "?l ! j = act (snd (?scmd ?srs) ! j) (snd ?scfg' ! j)" by simp moreover have "snd ?scfg' ! j = tps ! j" using shift_cfg_def assms(2) by (metis (no_types, lifting) Suc.IH True nth_append snd_conv) moreover have "snd (?scmd ?srs) ! j = (?srs ! j, Stay)" proof - have "snd (?scmd ?srs) = map (\h. (h, Stay)) (take d ?srs) @ snd ((M ! (fst ?scfg')) (drop d ?srs))" using scmd by simp then have "snd (?scmd ?srs) ! j = map (\h. (h, Stay)) (take d ?srs) ! j" using len_srs lenl True that - by (smt (z3) add.commute length_map length_take min_less_iff_conj nth_append) + by (smt (verit) add.commute length_map length_take min_less_iff_conj nth_append) then show ?thesis using len_srs True by simp qed moreover have "?r ! j = tps ! j" using True assms(2) by (simp add: nth_append) ultimately show ?thesis using len_s that lenl by (metis act_Stay) next case False have jle: "j < d + k" using lenl that by simp have jle': "j - d < k" using lenl that False by simp let ?at = "zip (snd (?scmd ?srs)) (snd ?scfg') ! j" have "?l ! j = act (fst ?at) (snd ?at)" using that by simp moreover have "fst ?at = snd (?scmd ?srs) ! j" using that by simp moreover have "snd ?at = snd ?scfg' ! j" using that by simp ultimately have "?l ! j = act (snd (?scmd ?srs) ! j) (snd ?scfg' ! j)" by simp moreover have "snd ?scfg' ! j = snd ?cfg' ! (j - d)" using shift_cfg_def assms(2) Suc False jle by (metis nth_append snd_conv) moreover have "snd (?scmd ?srs) ! j = snd (?cmd ?rs) ! (j - d)" proof - have "snd (?scmd ?srs) = map (\h. (h, Stay)) (take d ?srs) @ snd ((M ! (fst ?cfg')) (drop d ?srs))" using cmd by simp then have "snd (?scmd ?srs) ! j = snd ((M ! (fst ?cfg')) (drop d ?srs)) ! (j - d)" using len_srs lenl False that len_rs - by (smt (z3) Nat.add_diff_assoc add.right_neutral add_diff_cancel_left' append_take_drop_id + by (smt (verit) Nat.add_diff_assoc add.right_neutral add_diff_cancel_left' append_take_drop_id le_add1 length_append length_map nth_append srs_rs) then have "snd (?scmd ?srs) ! j = snd (?cmd ?rs) ! (j - d)" using srs_rs by simp then show ?thesis by simp qed moreover have "?r ! j = act (snd (?cmd ?rs) ! (j - d)) (snd ?cfg' ! (j - d))" proof - have "fst (execute M cfg0 t) < length M" using running fst prepend_tapes_def by simp then have len1: "length (snd (?cmd ?rs)) = k" using assms(1) len_rs turing_machine_def[of k G M] turing_commandD(1) by fastforce have "?r ! j = map (\(a, tp). act a tp) (zip (snd (?cmd ?rs)) (snd ?cfg')) ! (j - d)" using assms(2) False by (simp add: nth_append) also have "... = act (snd (?cmd ?rs) ! (j - d)) (snd ?cfg' ! (j - d))" using len1 len jle' by simp finally show ?thesis by simp qed ultimately show ?thesis by simp qed qed qed then show ?thesis using lhs rhs by simp qed qed lemma transforms_prepend_tapes: assumes "turing_machine k G M" and "length tps = d" and "length tps0 = k" and "transforms M tps0 t tps1" shows "transforms (prepend_tapes d M) (tps @ tps0) t (tps @ tps1)" proof - have "\t'\t. execute M (0, tps0) t' = (length M, tps1)" using assms(4) transforms_def transits_def by simp then have "\t'\t. execute (prepend_tapes d M) (shift_cfg tps (0, tps0)) t' = shift_cfg tps (length M, tps1)" using assms transforms_def transits_def execute_prepend_tapes shift_cfg_def by auto moreover have "length M = length (prepend_tapes d M)" using prepend_tapes_def by simp ultimately show ?thesis using shift_cfg_def transforms_def transitsI by auto qed end \ No newline at end of file diff --git a/thys/Cook_Levin/Combinations.thy b/thys/Cook_Levin/Combinations.thy --- a/thys/Cook_Levin/Combinations.thy +++ b/thys/Cook_Levin/Combinations.thy @@ -1,1009 +1,1009 @@ section \Combining Turing machines\label{s:tm-combining}\ theory Combinations imports Basics "HOL-Eisbach.Eisbach" begin text \ This section describes how we can combine Turing machines in the way of traditional control structures found in structured programming, namely sequences, branching, and iterating. This allows us to build complex Turing machines from simpler ones and analyze their behavior and running time. Thanks to some syntactic sugar the result may even look like a programming language, but it is really more like a ``construction kit'' than a ``true'' programming language with small and big step semantics or Hoare rules. Instead we will merely have some lemmas for proving @{const transforms} statements for the combined machines. The remaining sections of this chapter will provide us with concrete Turing machines to combine. \ subsection \Relocated machines\ text \ If we use a Turing machine $M$ as part of another TM and there are $q$ commands before $M$, then $M$'s target states will be off by $q$. This can be fixed by adding $q$ to all target states of all commands in $M$, an operation we call \emph{relocation}. \ definition relocate_cmd :: "nat \ command \ command" where "relocate_cmd q cmd rs \ (fst (cmd rs) + q, snd (cmd rs))" lemma relocate_cmd_head: "relocate_cmd q cmd rs [~] j = cmd rs [~] j" using relocate_cmd_def by simp lemma sem_relocate_cmd: "sem (relocate_cmd q cmd) cfg = (sem cmd cfg) <+=> q" proof- let ?cmd' = "relocate_cmd q cmd" let ?rs = "read (snd cfg)" have "?cmd' ?rs = (fst (cmd ?rs) + q, snd (cmd ?rs))" by (simp add: relocate_cmd_def) then show ?thesis using sem by (metis (no_types, lifting) fst_conv snd_conv) qed definition relocate :: "nat \ machine \ machine" where "relocate q M \ map (relocate_cmd q) M" lemma relocate: assumes "M' = relocate q M" and "i < length M" shows "(M' ! i) r = (fst ((M ! i) r) + q, snd ((M ! i) r))" using assms relocate_def relocate_cmd_def by simp lemma sem_relocate: assumes "M' = relocate q M" and "i < length M" shows "sem (M' ! i) cfg = sem (M ! i) cfg <+=> q" using assms relocate_def sem_relocate_cmd by simp lemma turing_command_relocate_cmd: assumes "turing_command k Q G cmd" shows "turing_command k (Q + q) G (relocate_cmd q cmd)" using assms relocate_cmd_def turing_commandD by auto lemma turing_command_relocate: assumes "M' = relocate q M" and "turing_machine k G M" and "i < length M" shows "turing_command k (length M + q) G (M' ! i)" proof fix gs :: "symbol list" assume gs: "length gs = k" have tc: "turing_command k (length M) G (M ! i)" using turing_machine_def assms(2,3) by simp show "length ([!!] (M' ! i) gs) = length gs" using gs turing_commandD(1)[OF tc] assms(1,3) relocate by simp show "(M' ! i) gs [.] 0 = gs ! 0" if "k > 0" using gs turing_commandD(3)[OF tc] assms(1,3) relocate that by simp show "[*] ((M' ! i) gs) \ length M + q" using gs turing_commandD(4)[OF tc] assms(1,3) relocate by simp show "(\i. i < length gs \ gs ! i < G) \ j < length gs \ (M' ! i) gs [.] j < G" for j using gs turing_commandD(2)[OF tc] assms(1,3) relocate by simp qed lemma length_relocate: "length (relocate q M) = length M" by (simp add: relocate_def) lemma relocate_jump_targets: assumes "turing_machine k G M" and "M' = relocate q M" and "i < length M" and "length rs = k" shows "fst ((M' ! i) rs) \ length M + q" using turing_machine_def relocate_def assms relocate by (metis turing_commandD(4) diff_add_inverse2 fst_conv le_diff_conv nth_mem) lemma relocate_zero: "relocate 0 M = M" unfolding relocate_def relocate_cmd_def by simp subsection \Sequences\ text \ To execute two Turing machines sequentially we concatenate the two machines, relocating the second one by the length of the first one. In this way the halting state of the first machine becomes the start state of the second machine. \ definition turing_machine_sequential :: "machine \ machine \ machine" (infixl ";;" 55) where "M1 ;; M2 \ M1 @ (relocate (length M1) M2)" text \ If the number of tapes and the alphabet size match, the concatenation of two Turing machines is again a Turing machine. \ lemma turing_machine_sequential_turing_machine [intro, simp]: assumes "turing_machine k G M1" and "turing_machine k G M2" shows "turing_machine k G (M1 ;; M2)" (is "turing_machine k G ?M") proof show 1: "k \ 2" using assms(1) turing_machine_def by simp show 2: "G \ 4" using assms(1) turing_machine_def by simp have len: "length ?M = length M1 + length M2" using relocate_def turing_machine_sequential_def by simp show 3: "turing_command k (length ?M) G (?M ! i)" if "i < length ?M" for i proof (cases "i < length M1") case True then show ?thesis using turing_machineD[OF assms(1)] turing_machine_sequential_def len turing_command_mono by (metis (no_types, lifting) le_add1 nth_append) next case False then have "i - (length M1) < length M2" (is "?i < length M2") using False that len by simp then have "turing_command k (length ?M) G ((relocate (length M1) M2) ! ?i)" using assms(2) turing_command_relocate len by (simp add: add.commute) moreover have "?M ! i = (relocate (length M1) M2) ! ?i" using False by (simp add: nth_append turing_machine_sequential_def) ultimately show ?thesis by simp qed qed lemma turing_machine_sequential_empty: "turing_machine_sequential [] M = M" unfolding turing_machine_sequential_def using relocate_zero by simp lemma turing_machine_sequential_nth: assumes "M = M1 ;; M2" and "p < length M2" shows "M ! (p + length M1) = relocate_cmd (length M1) (M2 ! p)" proof- let ?r = "relocate (length M1) M2" have "M = M1 @ ?r" using assms(1) turing_machine_sequential_def by simp then have "M ! (p + length M1) = ?r ! p" by (simp add: nth_append) then show ?thesis using assms(2) relocate_def by simp qed lemma turing_machine_sequential_nth': assumes "M = M1 ;; M2" and "length M1 \ q" and "q < length M" shows "M ! q = relocate_cmd (length M1) (M2 ! (q - length M1))" using assms turing_machine_sequential_nth length_relocate turing_machine_sequential_def by (metis (no_types, lifting) add.assoc diff_add length_append less_add_eq_less) lemma "length_turing_machine_sequential": "length (turing_machine_sequential M1 M2) = length M1 + length M2" using turing_machine_sequential_def relocate_def by simp lemma exe_relocate: "exe (M1 ;; M2) (cfg <+=> length M1) = (exe M2 cfg) <+=> length M1 " using sem_relocate_cmd sem_state_indep exe_def turing_machine_sequential_nth length_turing_machine_sequential by (smt (verit, ccfv_SIG) add.commute diff_add_inverse2 fst_conv le_add2 less_diff_conv2 snd_conv) lemma execute_pre_append: assumes "halts M1 cfg" and "fst cfg = 0" and "t \ running_time M1 cfg " shows "execute ((M0 ;; M1) @ M2) (cfg <+=> length M0) t = (execute M1 cfg t) <+=> length M0" using assms(3) proof (induction t) case 0 then show ?case by simp next case (Suc t) let ?l0 = "length M0" let ?M = "(M0 ;; M1) @ M2" let ?cfg_t = "execute ?M (cfg <+=> ?l0) t" let ?cfg1_t = "execute M1 cfg t" let ?cmd1 = "M1 ! (fst ?cfg1_t)" let ?cmd = "?M ! (fst ?cfg_t)" have *: "?cfg1_t <+=> ?l0 = ?cfg_t" using Suc by simp then have "fst (?cfg1_t <+=> ?l0) = fst ?cfg_t" by simp then have 2: "fst ?cfg1_t + ?l0 = fst ?cfg_t" by simp from * have sndeq: "snd ?cfg1_t = snd ?cfg_t" by (metis snd_conv) have "fst (execute M1 cfg t) \ length M1" using halts_impl_le_length assms(1) halts_def by blast moreover have "fst ?cfg1_t \ length M1" using Suc.prems running_time_def wellorder_Least_lemma(2) by fastforce ultimately have 1: "fst ?cfg1_t < length M1" by simp with 2 have "relocate_cmd ?l0 ?cmd1 = (M0 ;; M1) ! (fst ?cfg1_t + ?l0)" by (metis turing_machine_sequential_nth) then have "relocate_cmd ?l0 ?cmd1 = ?M ! (fst ?cfg1_t + ?l0)" by (simp add: "1" nth_append length_turing_machine_sequential) then have 3: "relocate_cmd ?l0 ?cmd1 = ?cmd" using 2 by simp with 1 have "execute M1 cfg (Suc t) = sem ?cmd1 ?cfg1_t" by (simp add: exe_def) then have "(execute M1 cfg (Suc t)) <+=> ?l0 = (sem ?cmd1 ?cfg1_t) <+=> ?l0" by simp then have "(execute M1 cfg (Suc t)) <+=> ?l0 = (sem (relocate_cmd ?l0 ?cmd1) ?cfg1_t)" using sem_relocate_cmd by simp then have rhs: "(execute M1 cfg (Suc t)) <+=> ?l0 = sem ?cmd ?cfg1_t" using 3 by simp have "execute ?M (cfg <+=> ?l0) (Suc t) = exe ?M ?cfg_t" by simp moreover have "fst ?cfg_t < length ?M" using 1 2 by (simp add: length_turing_machine_sequential) ultimately have lhs: "execute ?M (cfg <+=> ?l0) (Suc t) = sem ?cmd ?cfg_t" by (simp add: exe_lt_length) have "sem ?cmd ?cfg_t = sem ?cmd ?cfg1_t" using sem_state_indep sndeq by fastforce with lhs rhs show ?case by simp qed lemma transits_pre_append': assumes "transforms M1 tps t tps'" shows "transits ((M0 ;; M1) @ M2) (length M0, tps) t (length M0 + length M1, tps')" proof- let ?rt = "running_time M1 (0, tps)" let ?M = "(M0 ;; M1) @ M2" have "?rt \ t" using assms transforms_running_time by simp have "fst (execute M1 (0, tps) t) = length M1" using assms(1) execute_after_halting_ge halting_config_def transforms_halting_config transforms_running_time by (metis (no_types, opaque_lifting) fst_conv) then have *: "halts M1 (0, tps)" using halts_def by auto have "transits M1 (0, tps) t (length M1, tps')" using assms(1) transits_def by (simp add: transforms_def) then have "execute M1 (0, tps) ?rt = (length M1, tps')" using assms(1) halting_config_def transforms_halting_config by auto moreover have "execute ?M (length M0, tps) ?rt = (execute M1 (0, tps) ?rt) <+=> length M0" using execute_pre_append * by auto ultimately have "execute ?M (length M0, tps) ?rt = (length M1, tps') <+=> length M0" by simp then have "execute ?M (length M0, tps) ?rt = (length M0 + length M1, tps')" by simp then show ?thesis using \?rt \ t\ transits_def by blast qed corollary transits_prepend: assumes "transforms M1 tps t tps'" shows "transits (M0 ;; M1) (length M0, tps) t (length M0 + length M1, tps')" using transits_pre_append' assms by (metis append_Nil2) corollary transits_append: assumes "transforms M1 tps t tps'" shows "transits (M1 @ M2) (0, tps) t (length M1, tps')" using transits_pre_append' turing_machine_sequential_empty assms by (metis gen_length_def length_code list.size(3)) corollary execute_append: assumes "fst cfg = 0" and "halts M1 cfg" and "t \ running_time M1 cfg" shows "execute (M1 @ M2) cfg t = execute M1 cfg t" using assms execute_pre_append turing_machine_sequential_empty by (metis add.right_neutral list.size(3) prod.collapse) lemma execute_sequential: assumes "execute M1 cfg1 t1 = cfg1'" and "fst cfg1 = 0" and "t1 = running_time M1 cfg1" and "execute M2 cfg2 t2 = cfg2'" and "cfg1' = cfg2 <+=> length M1" and "halts M1 cfg1" shows "execute (M1 ;; M2) cfg1 (t1 + t2) = cfg2' <+=> length M1" proof- let ?M = "M1 ;; M2" have 1: "execute ?M cfg1 t1 = cfg1'" using execute_append assms turing_machine_sequential_def by simp then have 2: "execute ?M cfg1 (t1 + t2) = execute ?M cfg1' t2" using execute_additive by simp have "execute M2 cfg2 t2 = cfg2' \ execute ?M cfg1' t2 = cfg2' <+=> length M1" for t2 proof (induction t2 arbitrary: cfg2') case 0 then show ?case using 1 assms(5) by simp next case (Suc t2) let ?cfg = "execute M2 cfg2 t2" have "execute ?M cfg1' (Suc t2) = exe ?M (execute ?M cfg1' t2)" by simp then have "execute ?M cfg1' (Suc t2) = exe ?M (?cfg <+=> length M1)" using Suc by simp moreover have "execute M2 cfg2 (Suc t2) = exe M2 ?cfg" by simp ultimately show ?case using exe_relocate Suc.prems by metis qed then show ?thesis using assms(4) 2 by simp qed text \ The semantics and running time of the @{text ";;"} operator: \ lemma transforms_turing_machine_sequential: assumes "transforms M1 tps1 t1 tps2" and "transforms M2 tps2 t2 tps3" shows "transforms (M1 ;; M2) tps1 (t1 + t2) tps3" proof- let ?M = "M1 ;; M2" let ?cfg1 = "(0, tps1)" let ?cfg1' = "(length M1, tps2)" let ?t1 = "running_time M1 ?cfg1" let ?cfg2 = "(0, tps2)" let ?cfg2' = "(length M2, tps3)" let ?t2 = "running_time M2 ?cfg2" have "fst (execute M1 ?cfg1 ?t1) = length M1" using assms(1) halting_config_def transforms_halting_config by (metis fst_conv) then have *: "halts M1 ?cfg1" using halts_def by auto have "execute M1 ?cfg1 ?t1 = ?cfg1'" using assms(1) halting_config_def transforms_halting_config by auto moreover have "fst ?cfg1 = 0" by simp moreover have "execute M2 ?cfg2 ?t2 = ?cfg2'" using assms(2) halting_config_def transforms_halting_config by auto moreover have "?cfg1' = ?cfg2 <+=> length M1" by simp ultimately have "execute ?M ?cfg1 (?t1 + ?t2) = ?cfg2' <+=> length M1" using execute_sequential * by blast then have "execute ?M ?cfg1 (?t1 + ?t2) = (length ?M, tps3)" by (simp add: length_turing_machine_sequential) then have "transits ?M ?cfg1 (?t1 + ?t2) (length ?M, tps3)" using transits_def by blast moreover have "?t1 + ?t2 \ t1 + t2" using add_le_mono assms(1,2) transforms_running_time by blast ultimately have "transits ?M ?cfg1 (t1 + t2) (length ?M, tps3)" using transits_monotone by simp then show ?thesis using transforms_def by simp qed corollary transforms_tm_sequentialI: assumes "transforms M1 tps1 t1 tps2" and "transforms M2 tps2 t2 tps3" and "t12 = t1 + t2" shows "transforms (M1 ;; M2) tps1 t12 tps3" using assms transforms_turing_machine_sequential by simp subsection \Branches\ text \ A branching Turing machine consists of a condition and two Turing machines, one for each of the branches. The condition can be any predicate over the list of symbols read from the tapes. The branching TM thus needs to perform conditional jumps, for which we will have the following command: \ definition cmd_jump :: "(symbol list \ bool) \ state \ state \ command" where "cmd_jump cond q1 q2 rs \ (if cond rs then q1 else q2, map (\r. (r, Stay)) rs)" lemma turing_command_jump_1: assumes "q1 \ q2" and "k > 0" shows "turing_command k q2 G (cmd_jump cond q1 q2)" using assms cmd_jump_def turing_commandI by simp lemma turing_command_jump_2: assumes "q2 \ q1" and "k > 0" shows "turing_command k q1 G (cmd_jump cond q1 q2)" using assms cmd_jump_def turing_commandI by simp lemma sem_jump_snd: "snd (sem (cmd_jump cond q1 q2) cfg) = snd cfg" proof- let ?k = "||cfg||" let ?cmd = "cmd_jump cond q1 q2" let ?cfg' = "sem ?cmd cfg" let ?rs = "read (snd cfg)" have 1: "proper_command ?k ?cmd" using cmd_jump_def by simp then have len: "||?cfg'|| = ||cfg||" using sem_num_tapes_raw by simp have "snd ?cfg' ! i = act (snd (?cmd ?rs) ! i) (snd cfg ! i)" if "i < ||cfg||" for i using sem_snd 1 that by simp moreover have "snd (?cmd ?rs) ! i = (?rs!i, Stay)" if "i < ||cfg||" for i using cmd_jump_def by (simp add: read_length that) ultimately have "snd ?cfg' ! i = act (?rs ! i, Stay) (snd cfg ! i)" if "i < ||cfg||" for i using that by simp then have "snd ?cfg' ! i = snd cfg ! i" if "i < ||cfg||" for i using that act_Stay by simp then show "snd ?cfg' = snd cfg" using len nth_equalityI by force qed lemma sem_jump_fst1: assumes "cond (read (snd cfg))" shows "fst (sem (cmd_jump cond q1 q2) cfg) = q1" using cmd_jump_def sem assms by simp lemma sem_jump_fst2: assumes "\ cond (read (snd cfg))" shows "fst (sem (cmd_jump cond q1 q2) cfg) = q2" using cmd_jump_def sem assms by simp lemma sem_jump: "sem (cmd_jump cond q1 q2) cfg = (if cond (config_read cfg) then q1 else q2, snd cfg)" using sem_def sem_jump_snd cmd_jump_def by simp lemma transits_jump: "transits [cmd_jump cond q1 q2] (0, tps) 1 (if cond (read tps) then q1 else q2, tps)" using transits_def sem_jump exe_def by auto definition turing_machine_branch :: "(symbol list \ bool) \ machine \ machine \ machine" ("IF _ THEN _ ELSE _ ENDIF" 60) where "IF cond THEN M1 ELSE M2 ENDIF \ [cmd_jump cond 1 (length M1 + 2)] @ (relocate 1 M1) @ [cmd_jump (\_. True) (length M1 + length M2 + 2) 0] @ (relocate (length M1 + 2) M2)" lemma turing_machine_branch_len: "length (IF cond THEN M1 ELSE M2 ENDIF) = length M1 + length M2 + 2" unfolding turing_machine_branch_def by (simp add: relocate_def) text \ If the Turing machines for both branches have the same number of tapes and the same alphabet size, the branching machine is a Turing machine, too. \ lemma turing_machine_branch_turing_machine: assumes "turing_machine k G M1" and "turing_machine k G M2" shows "turing_machine k G (IF cond THEN M1 ELSE M2 ENDIF)" (is "turing_machine _ _ ?M") proof show "k \ 2" using assms(2) turing_machine_def by simp show "G \ 4" using assms(2) turing_machine_def by simp let ?C1 = "[cmd_jump cond 1 (length M1 + 2)]" let ?C2 = "relocate 1 M1" let ?C3 = "[cmd_jump (\_. True) (length M1 + length M2 + 2) 0]" let ?C4 = "relocate (length M1 + 2) M2" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_branch_def by simp have len: "length ?M = length M1 + length M2 + 2" using turing_machine_branch_len by simp have "k > 0" using `k \ 2` by simp show "turing_command k (length ?M) G (?M ! i)" if "i < length ?M" for i proof - consider "i < length ?C1" | "length ?C1 \ i \ i < length (?C1 @ ?C2)" | "length (?C1 @ ?C2) \ i \ i < length (?C1 @ ?C2 @ ?C3)" | "length (?C1 @ ?C2 @ ?C3) \ i \ i < length ?M" using `i < length ?M` by linarith then show ?thesis proof (cases) case 1 then have "turing_command k (length M1 + 2) G (?C1 ! i)" using turing_command_jump_1 `k > 0` by simp then have "turing_command k (length ?M) G (?C1 ! i)" using turing_command_mono len by simp then show ?thesis using parts 1 by simp next case 2 then have "i - length ?C1 < length ?C2" by auto then have "turing_command k (length M1 + 1) G (?C2 ! (i - length ?C1))" using turing_command_relocate assms length_relocate by metis then have "turing_command k (length ?M) G (?C2 ! (i - length ?C1))" using turing_command_mono len by simp then have "turing_command k (length ?M) G ((?C1 @ ?C2) ! i)" using 2 by simp then show ?thesis using parts 2 by (metis (no_types, lifting) append.assoc nth_append) next case 3 then have "turing_command k (length ?M) G (?C3 ! (i - length (?C1 @ ?C2)))" using turing_command_jump_2 len `k > 0` by simp then have "turing_command k (length ?M) G ((?C1 @ ?C2 @ ?C3) ! i)" using 3 by (metis (no_types, lifting) append.assoc diff_is_0_eq' less_numeral_extra(3) nth_append zero_less_diff) then show ?thesis using parts 3 by (smt (verit, best) append.assoc nth_append) next case 4 then have "i - length (?C1 @ ?C2 @ ?C3) < length ?C4" by (simp add: len less_diff_conv2 length_relocate) then have "turing_command k (length ?M) G (?C4 ! (i - length (?C1 @ ?C2 @ ?C3)))" using turing_command_relocate assms by (smt (verit, ccfv_SIG) add.assoc add.commute len length_relocate) then show ?thesis using parts 4 by (metis (no_types, lifting) append.assoc le_simps(3) not_less_eq_eq nth_append) qed qed qed text \ If the condition is true, the branching TM executes $M_1$ and requires two extra steps: one for evaluating the condition and one for the unconditional jump to the halting state. \ lemma transforms_branch_true: assumes "transforms M1 tps t tps'" and "cond (read tps)" shows "transforms (IF cond THEN M1 ELSE M2 ENDIF) tps (t + 2) tps'" (is "transforms ?M _ _ _") proof- let ?C1 = "[cmd_jump cond 1 (length M1 + 2)]" let ?C2 = "relocate 1 M1" let ?C3 = "[cmd_jump (\_. True) (length M1 + length M2 + 2) 0]" let ?C4 = "relocate (length M1 + 2) M2" let ?C34 = "?C3 @ ?C4" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_branch_def by simp then have "?M = ?C1 @ ?C2 @ ?C34" by simp moreover have "?C1 @ ?C2 = ?C1 ;; M1" using turing_machine_sequential_def by simp ultimately have parts2: "?M = (?C1 ;; M1) @ ?C34" by simp have "execute ?M (0, tps) 1 = exe ?M (0, tps)" by simp also have "... = sem (?M ! 0) (0, tps)" using exe_def by (simp add: turing_machine_branch_len) also have "... = sem (?C1 ! 0) (0, tps)" by (simp add: parts) also have "... = sem (cmd_jump cond 1 (length M1 + 2)) (0, tps)" by simp also have "... = (1, tps)" using sem_jump by (simp add: assms(2)) finally have "execute ?M (0, tps) 1 = (1, tps)" . then have phase1: "transits ?M (0, tps) 1 (1, tps)" using transits_def by auto have "length ?C1 = 1" by simp moreover have "transits ((?C1 ;; M1) @ ?C34) (length ?C1, tps) t (length ?C1 + length M1, tps')" using transits_pre_append' assms by blast ultimately have "transits ?M (1, tps) t (1 + length M1, tps')" using parts2 by simp with phase1 have "transits ?M (0, tps) (t + 1) (1 + length M1, tps')" using transits_additive by fastforce then have phase2: "transits ?M (0, tps) (t + 1) (length (?C1 @ ?C2), tps')" by (simp add: relocate_def) let ?cfg = "(length (?C1 @ ?C2), tps')" have *: "?M ! (length (?C1 @ ?C2)) = ?C3 ! 0" using parts by simp then have "execute ?M ?cfg 1 = exe ?M ?cfg" by simp also have "... = sem (cmd_jump (\_. True) (length M1 + length M2 + 2) 0) ?cfg" using exe_def relocate_def turing_machine_branch_len * by (simp add: Suc_le_lessD) also have "... = (length M1 + length M2 + 2, snd ?cfg)" using sem_jump by simp also have "... = (length ?M, tps')" by (simp add: turing_machine_branch_len) finally have "execute ?M ?cfg 1 = (length ?M, tps')" . then have "transits ?M ?cfg 1 (length ?M, tps')" using transits_def by auto with phase2 have "transits ?M (0, tps) (t + 2) (length ?M, tps')" using transits_additive by fastforce then show ?thesis by (simp add: transforms_def) qed text \ If the condition is false, the branching TM executes $M_2$ and requires one extra step to evaluate the condition. \ lemma transforms_branch_false: assumes "transforms M2 tps t tps'" and "\ cond (read tps)" shows "transforms (IF cond THEN M1 ELSE M2 ENDIF) tps (t + 1) tps'" (is "transforms ?M _ _ _") proof- let ?C1 = "[cmd_jump cond 1 (length M1 + 2)]" let ?C2 = "relocate 1 M1" let ?C3 = "[cmd_jump (\_. True) (length M1 + length M2 + 2) 0]" let ?C4 = "relocate (length M1 + 2) M2" let ?C123 = "?C1 @ ?C2 @ ?C3" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_branch_def by simp moreover have len123: "length ?C123 = length M1 + 2" by (simp add: length_relocate) ultimately have seq: "?M = ?C123 ;; M2" by (simp add: turing_machine_sequential_def) have "execute ?M (0, tps) 1 = exe ?M (0, tps)" by simp also have "... = sem (?M ! 0) (0, tps)" using exe_def by (simp add: turing_machine_branch_len) also have "... = sem (cmd_jump cond 1 (length M1 + 2)) (0, tps)" by (simp add: parts) also have "... = (length M1 + 2, tps)" using assms(2) sem_jump by simp also have "... = (length ?C123, tps)" using len123 by simp finally have "execute ?M (0, tps) 1 = (length ?C123, tps)" . then have phase1: "transits ?M (0, tps) 1 (length ?C123, tps)" using transits_def by blast have "?M ! (length ?C123) = ?C4 ! 0" by (simp add: nth_append parts length_relocate) have "transits (?C123 ;; M2) (length ?C123, tps) t (length ?C123 + length M2, tps')" using transits_prepend assms by blast then have "transits ?M (length ?C123, tps) t (length ?M, tps')" by (simp add: seq length_turing_machine_sequential) with phase1 have "transits ?M (0, tps) (t + 1) (length ?M, tps')" using transits_additive by fastforce then show ?thesis using transforms_def by simp qed text \ The behavior and running time of the branching Turing machine: \ lemma transforms_branch_full: assumes "c \ transforms M1 tps tT tpsT" and "\ c \ transforms M2 tps tF tpsF" and "c \ tT + 2 \ t" and "\ c \ tF + 1 \ t" and "c = cond (read tps)" and "tps' = (if c then tpsT else tpsF)" shows "transforms (IF cond THEN M1 ELSE M2 ENDIF) tps t tps'" proof - have "transforms (IF cond THEN M1 ELSE M2 ENDIF) tps (if c then tT + 2 else tF + 1) (if c then tpsT else tpsF)" using assms(1,2,5) transforms_branch_true transforms_branch_false by simp moreover have "(if c then tT + 2 else tF + 1) \ t" using assms(3,4) by simp ultimately show ?thesis using transforms_monotone assms(6) by presburger qed corollary transforms_branchI: assumes "cond (read tps) \ transforms M1 tps tT tpsT" and "\ cond (read tps) \ transforms M2 tps tF tpsF" and "cond (read tps) \ tT + 2 \ t" and "\ cond (read tps) \ tF + 1 \ t" and "cond (read tps) \ tps' = tpsT" and "\ cond (read tps) \ tps' = tpsF" shows "transforms (IF cond THEN M1 ELSE M2 ENDIF) tps t tps'" - using transforms_branch_full assms by (smt (z3)) + by (rule transforms_branch_full) (use assms in auto) subsection \Loops\ text \ The loops are while loops consisting of a head and a body. The body is a Turing machine that is executed in every iteration as long as the condition in the head of the loop evaluates to true. The condition is of the same form as in branching TMs, namely a predicate over the symbols read from the tapes. Sometimes this is not expressive enough, and so we allow a Turing machine as part of the loop head that is run prior to evaluating the condition. In most cases, however, this TM is empty. \ definition turing_machine_loop :: "machine \ (symbol list \ bool) \ machine \ machine" ("WHILE _ ; _ DO _ DONE" 60) where "WHILE M1 ; cond DO M2 DONE \ M1 @ [cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)] @ (relocate (length M1 + 1) M2) @ [cmd_jump (\_. True) 0 0]" text \ Intuitively the Turing machine @{term "WHILE M1 ; cond DO M2 DONE"} first executes @{term M1} and then checks the condition @{term cond}. If it is true, it executes @{term M2} and jumps back to the start state; otherwise it jumps to the halting state. \ lemma turing_machine_loop_len: "length (WHILE M1 ; cond DO M2 DONE) = length M1 + length M2 + 2" unfolding turing_machine_loop_def by (simp add: relocate_def) text \ If both Turing machines have the same number of tapes and alphabet size, then so does the looping Turing machine. \ lemma turing_machine_loop_turing_machine: assumes "turing_machine k G M1" and "turing_machine k G M2" shows "turing_machine k G (WHILE M1 ; cond DO M2 DONE)" (is "turing_machine _ _ ?M") proof show 1: "k \ 2" using assms(1) turing_machine_def by simp show 2: "G \ 4" using assms(1) turing_machine_def by simp let ?C1 = "M1" let ?C2 = "[cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)]" let ?C3 = "relocate (length M1 + 1) M2" let ?C4 = "[cmd_jump (\_. True) 0 0]" let ?C34 = "?C3 @ ?C4" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_loop_def by simp have len: "length ?M = length M1 + length M2 + 2" using turing_machine_loop_len by simp have "k > 0" using `k \ 2` by simp show "turing_command k (length ?M) G (?M ! i)" if "i < length ?M" for i proof - consider "i < length ?C1" | "length ?C1 \ i \ i < length (?C1 @ ?C2)" | "length (?C1 @ ?C2) \ i \ i < length (?C1 @ ?C2 @ ?C3)" | "length (?C1 @ ?C2 @ ?C3) \ i \ i < length ?M" using `i < length ?M` by linarith then show ?thesis proof (cases) case 1 then have "turing_command k (length M1) G (?C1 ! i)" using turing_machineD(3) assms by simp then have "turing_command k (length ?M) G (?C1 ! i)" using turing_command_mono len by simp then show ?thesis using parts 1 by (simp add: nth_append) next case 2 then have "turing_command k (length M1 + length M2 + 2) G (?C2 ! (i - length ?C1))" using turing_command_jump_1 `0 < k` by simp then have "turing_command k (length ?M) G (?C2 ! (i - length ?C1))" using len by simp then have "turing_command k (length ?M) G ((?C1 @ ?C2) ! i)" using "2" le_add_diff_inverse by (simp add: nth_append) then show ?thesis using 2 parts by (simp add: nth_append) next case 3 then have "turing_command k (length M2 + (length M1 + 1)) G (?C3 ! (i - length (?C1 @ ?C2)))" using turing_command_relocate length_relocate assms(2) by (smt (verit, best) add.commute add.left_commute add_less_cancel_left le_add_diff_inverse length_append) then have "turing_command k (length ?M) G (?C3 ! (i - length (?C1 @ ?C2)))" using turing_command_mono len by simp then have "turing_command k (length ?M) G ((?C1 @ ?C2 @ ?C3) ! i)" using 3 by (simp add: nth_append) then show ?thesis using parts 3 by (smt (verit) append.assoc nth_append) next case 4 then have "turing_command k 0 G (?C4 ! (i - length (?C1 @ ?C2 @ ?C3)))" using turing_command_jump_1 `0 < k` len length_relocate by simp then have "turing_command k (length ?M) G (?C4 ! (i - length (?C1 @ ?C2 @ ?C3)))" using turing_command_mono by blast then show ?thesis using parts 4 by (metis (no_types, lifting) append.assoc nth_append verit_comp_simplify1(3)) qed qed qed lemma transits_turing_machine_loop_cond_false: assumes "transforms M1 tps t1 tps'" and "\ cond (read tps')" shows "transits (WHILE M1 ; cond DO M2 DONE) (0, tps) (t1 + 1) (length M1 + length M2 + 2, tps')" (is "transits ?M _ _ _") proof- let ?C1 = "M1" let ?C2 = "[cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)]" let ?C3 = "relocate (length M1 + 1) M2" let ?C4 = "[cmd_jump (\_. True) 0 0]" let ?C34 = "?C3 @ ?C4" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_loop_def by simp then have "?M = ?C1 @ (?C2 @ ?C3 @ ?C4)" by simp then have "transits ?M (0, tps) t1 (length ?C1, tps')" using assms transits_append by simp moreover have "transits ?M (length M1, tps') 1 (length M1 + length M2 + 2, tps')" proof- have *: "?M ! (length ?C1) = cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)" using turing_machine_loop_def by simp have "execute ?M (length ?C1, tps') 1 = exe ?M (length ?C1, tps')" by simp also have "... = sem (?M ! (length ?C1)) (length ?C1, tps')" by (simp add: exe_lt_length turing_machine_loop_len) also have "... = sem (cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)) (length ?C1, tps')" using * by simp also have "... = (length M1 + length M2 + 2, tps')" using sem_jump assms(2) by simp finally have "execute ?M (length ?C1, tps') 1 = (length M1 + length M2 + 2, tps')" . then show ?thesis using transits_def by auto qed ultimately show "transits ?M (0, tps) (t1 + 1) (length M1 + length M2 + 2, tps')" using transits_additive by blast qed lemma transits_turing_machine_loop_cond_true: assumes "transforms M1 tps t1 tps'" and "transforms M2 tps' t2 tps''" and "cond (read tps')" shows "transits (WHILE M1 ; cond DO M2 DONE) (0, tps) (t1 + t2 + 2) (0, tps'')" (is "transits ?M _ _ _") proof- let ?C1 = "M1" let ?C2 = "[cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)]" let ?C3 = "relocate (length M1 + 1) M2" let ?C4 = "[cmd_jump (\_. True) 0 0]" let ?C34 = "?C3 @ ?C4" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_loop_def by simp then have "?M = ?C1 @ (?C2 @ ?C3 @ ?C4)" by simp then have "transits ?M (0, tps) t1 (length ?C1, tps')" using assms(1,3) transits_append by simp moreover have "transits ?M (length ?C1, tps') 1 (length ?C1 + 1, tps')" proof- have *: "?M ! (length ?C1) = cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)" using turing_machine_loop_def by simp have "execute ?M (length ?C1, tps') 1 = exe ?M (length ?C1, tps')" by simp also have "... = sem (?M ! (length ?C1)) (length ?C1, tps')" by (simp add: exe_lt_length turing_machine_loop_len) also have "... = sem (cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)) (length ?C1, tps')" using * by simp also have "... = (length M1 + 1, tps')" using sem_jump assms(3) by simp finally have "execute ?M (length ?C1, tps') 1 = (length M1 + 1, tps')" . then show ?thesis using transits_def by auto qed ultimately have "transits ?M (0, tps) (t1 + 1) (length M1 + 1, tps')" using transits_additive by blast moreover have "transits ?M (length M1 + 1, tps') t2 (length M1 + length M2 + 1, tps'')" proof- have "?M = ((?C1 @ ?C2) ;; M2) @ ?C4" by (simp add: parts turing_machine_sequential_def) moreover have "length (?C1 @ ?C2) = length M1 + 1" by simp ultimately have "transits ?M (length M1 + 1, tps') t2 (length M1 + 1 + length M2, tps'')" using assms transits_pre_append' by metis then show ?thesis by simp qed ultimately have "transits ?M (0, tps) (t1 + t2 + 1) (length M1 + length M2 + 1, tps'')" using transits_additive by fastforce moreover have "transits ?M (length M1 + length M2 + 1, tps'') 1 (0, tps'')" proof- have *: "?M ! (length M1 + length M2 + 1) = cmd_jump (\_. True) 0 0" by (simp add: nth_append parts length_relocate) have "execute ?M (length M1 + length M2 + 1, tps'') 1 = exe ?M (length M1 + length M2 + 1, tps'')" by simp also have "... = sem (?M ! (length M1 + length M2 + 1)) (length M1 + length M2 + 1, tps'')" by (simp add: exe_lt_length turing_machine_loop_len) also have "... = sem (cmd_jump (\_. True) 0 0) (length M1 + length M2 + 1, tps'')" using * by simp also have "... = (0, tps'')" using sem_jump by simp finally have "execute ?M (length M1 + length M2 + 1, tps'') 1 = (0, tps'')" . then show ?thesis using transits_def by auto qed ultimately show "transits ?M (0, tps) (t1 + t2 + 2) (0, tps'')" using transits_additive by fastforce qed text \ In this article we will only encounter while loops that are in fact for loops, that is, where the number of iterations is known beforehand. Moreover, using the same time bound for every iteration will lead to a good enough upper bound for the entire loop. The @{const transforms} rule for a loop with $m$ iterations where the running time of both TMs is bounded by a constant: \ lemma transforms_loop_for: fixes m t1 t2 :: nat and M1 M2 :: machine and tps :: "nat \ tape list" and tps' :: "nat \ tape list" assumes "\i. i \ m \ transforms M1 (tps i) t1 (tps' i)" assumes "\i. i < m \ transforms M2 (tps' i) t2 (tps (Suc i))" and "\i. i < m \ cond (read (tps' i))" and "\ cond (read (tps' m))" assumes "ttt \ m * (t1 + t2 + 2) + t1 + 1" shows "transforms (WHILE M1 ; cond DO M2 DONE) (tps 0) ttt (tps' m)" proof - define M where "M = WHILE M1 ; cond DO M2 DONE" define t where "t = t1 + t2 + 2" have "transits M (0, tps 0) (i * t) (0, tps i)" if "i \ m" for i using that proof (induction i) case 0 then show ?case using transits_def by simp next case (Suc i) then have "i < m" by simp then have "transits M (0, tps i) t (0, tps (Suc i))" using M_def t_def assms transits_turing_machine_loop_cond_true by (metis le_eq_less_or_eq) moreover have "transits M (0, tps 0) (i * t) (0, tps i)" using Suc by simp ultimately have "transits M (0, tps 0) (i * t + t) (0, tps (Suc i))" using transits_additive by simp then show ?case by (metis add.commute mult_Suc) qed then have "transits M (0, tps 0) (m * t) (0, tps m)" by simp moreover have "transits M (0, tps m) (t1 + 1) (length M1 + length M2 + 2, tps' m)" using assms(1,4) transits_turing_machine_loop_cond_false M_def by simp ultimately have "transits M (0, tps 0) (m * t + t1 + 1) (length M1 + length M2 + 2, tps' m)" using transits_additive by fastforce then show ?thesis using transforms_def turing_machine_loop_len M_def assms(5) t_def transits_monotone by auto qed text \ The rule becomes even simpler in the common case in which the Turing machine in the loop head is empty. \ lemma transforms_loop_simple: fixes m t :: nat and M :: machine and tps :: "nat \ tape list" assumes "\i. i < m \ transforms M (tps i) t (tps (Suc i))" and "\i. i < m \ cond (read (tps i))" and "\ cond (read (tps m))" assumes "ttt \ m * (t + 2) + 1" shows "transforms (WHILE [] ; cond DO M DONE) (tps 0) ttt (tps m)" using transforms_loop_for[where ?tps'=tps and ?cond=cond and ?t1.0=0, OF _ assms(1) _ assms(3)] transforms_Nil assms(2,4) by simp subsection \A proof method\ text \ Statements about the behavior and running time of Turing machines, expressed via the predicate @{const transforms}, are the most common ones we are going to prove. The following proof method applies introduction rules for this predicate. These rules are either the ones we proved for the control structures (sequence, branching, loop) or the ones describing the semantics of concrete Turing machines. The rules of the second kind are collected in the attribute @{text transforms_intros}. Applying such a rule usually leaves three kinds of goals: some simple ones requiring only instantiation of schematic variables; one for the equality of two tape lists; and one for the time bound. For the last two goals the proof method offers parameters @{term tps} and @{term time}, respectively. I have to admit that most of the details of the proof method were determined by trial and error. \null \ named_theorems transforms_intros declare transforms_Nil [transforms_intros] method tform uses tps time declares transforms_intros = ( ((rule transforms_tm_sequentialI)+ | rule transforms_branchI | rule transforms_loop_simple | rule transforms_loop_for | rule transforms_intros) ; (rule transforms_intros)? ; (simp only:; fail)? ; ((use tps in simp; fail) | (use time in simp; fail))? ; (match conclusion in "left = right" for left right :: "tape list" \ \(fastforce simp add: tps list_update_swap; fail)\)? ; (match conclusion in "left = right" for left right :: nat \ \(use time in simp; fail)?\)? ) text \ These lemmas are sometimes helpful for proving the equality of tape lists: \ lemma list_update_swap_less: "i' < i \ ys[i := x, i' := x'] = ys[i' := x', i := x]" by (simp add: list_update_swap) lemma nth_list_update_neq': "j \ i \ xs[i := x] ! j = xs ! j" by simp end diff --git a/thys/Cook_Levin/Composing.thy b/thys/Cook_Levin/Composing.thy --- a/thys/Cook_Levin/Composing.thy +++ b/thys/Cook_Levin/Composing.thy @@ -1,786 +1,786 @@ section \Composing functions\label{s:tm-composing}\ theory Composing imports Elementary begin text \ For a Turing machine $M_1$ computing $f_1$ in time $T_1$ and a TM $M_2$ computing $f_2$ in time $T_2$ there is a TM $M$ computing $f_2 \circ f_1$ in time $O(T_1(n) + \max_{m \le T_1(n)} T_2(m))$. If $T_1$ is monotone the time bound is $O(T_1 + T_2 \circ T_1)$; if $T_1$ and $T_2$ are polynomially bounded the running-time of $M$ is polynomially bounded, too. The Turing machines $M_1$ and $M_2$ can have both a different alphabet and number of tapes, so generally they cannot be composed by the $;;$ operator. To get around this we enlarge the alphabet and prepend and append tapes, so $M$ has as many tapes as $M_1$ and $M_2$ combined. The following function returns the combined Turing machine $M$. \null \ definition "compose_machines k1 G1 M1 k2 G2 M2 \ enlarged G1 (append_tapes k1 (k1 + k2) M1) ;; tm_start 1 ;; tm_cp_until 1 k1 {\} ;; tm_erase_cr 1 ;; tm_start k1 ;; prepend_tapes k1 (enlarged G2 M2) ;; tm_cr (k1 + 1) ;; tm_cp_until (k1 + 1) 1 {\}" locale compose = fixes k1 k2 G1 G2 :: nat and M1 M2 :: machine and T1 T2 :: "nat \ nat" and f1 f2 :: "string \ string" assumes tm_M1: "turing_machine k1 G1 M1" and tm_M2: "turing_machine k2 G2 M2" and computes1: "computes_in_time k1 M1 f1 T1" and computes2: "computes_in_time k2 M2 f2 T2" begin definition "tm1 \ enlarged G1 (append_tapes k1 (k1 + k2) M1)" definition "tm2 \ tm1 ;; tm_start 1" definition "tm3 \ tm2 ;; tm_cp_until 1 k1 {\}" definition "tm4 \ tm3 ;; tm_erase_cr 1" definition "tm5 \ tm4 ;; tm_start k1" definition "tm56 \ prepend_tapes k1 (enlarged G2 M2)" definition "tm6 \ tm5 ;; tm56" definition "tm7 \ tm6 ;; tm_cr (k1 + 1)" definition "tm8 \ tm7 ;; tm_cp_until (k1 + 1) 1 {\}" definition G :: nat where "G \ max G1 G2" lemma G1: "G1 \ G" and G2: "G2 \ G" using G_def by simp_all lemma k_ge: "k1 \ 2" "k2 \ 2" using tm_M1 tm_M2 turing_machine_def by simp_all lemma tm1_tm: "turing_machine (k1 + k2) G tm1" unfolding tm1_def using turing_machine_enlarged append_tapes_tm tm_M1 G1 by simp lemma tm2_tm: "turing_machine (k1 + k2) G tm2" unfolding tm2_def using tm1_tm tm_start_tm turing_machine_def by blast lemma tm3_tm: "turing_machine (k1 + k2) G tm3" unfolding tm3_def using tm2_tm tm_cp_until_tm turing_machine_def k_ge turing_machine_sequential_turing_machine by (metis add_leD1 less_add_same_cancel1 less_le_trans less_numeral_extra(1) nat_1_add_1) lemma tm4_tm: "turing_machine (k1 + k2) G tm4" unfolding tm4_def using tm3_tm tm_erase_cr_tm turing_machine_def turing_machine_sequential_turing_machine by (metis Suc_1 Suc_le_lessD tm_erase_cr_tm zero_less_one) lemma tm5_tm: "turing_machine (k1 + k2) G tm5" unfolding tm5_def using tm4_tm tm_start_tm turing_machine_def turing_machine_sequential_turing_machine by auto lemma tm6_tm: "turing_machine (k1 + k2) G tm6" unfolding tm6_def using tm5_tm tm56_def turing_machine_enlarged prepend_tapes_tm tm_M2 G2 by simp lemma tm7_tm: "turing_machine (k1 + k2) G tm7" unfolding tm7_def using tm6_tm tm_cr_tm turing_machine_def by blast lemma tm8_tm: "turing_machine (k1 + k2) G tm8" unfolding tm8_def using tm7_tm tm_cp_until_tm turing_machine_def turing_machine_sequential_turing_machine k_ge(2) by (metis add.commute add_less_cancel_right add_strict_increasing nat_1_add_1 verit_comp_simplify1(3) zero_less_one) context fixes x :: string begin definition "zs \ string_to_symbols x" lemma bit_symbols_zs: "bit_symbols zs" using zs_def by simp abbreviation "n \ length x" lemma length_zs [simp]: "length zs = n" using zs_def by simp definition "tps0 \ snd (start_config (k1 + k2) zs)" definition tps1a :: "tape list" where "tps1a \ SOME tps. tps ::: 1 = string_to_contents (f1 x) \ transforms M1 (snd (start_config k1 (string_to_symbols x))) (T1 n) tps" lemma tps1a_aux: "tps1a ::: 1 = string_to_contents (f1 x)" "transforms M1 (snd (start_config k1 (string_to_symbols x))) (T1 n) tps1a" using tps1a_def someI_ex[OF computes_in_timeD[OF computes1, of x]] by simp_all lemma tps1a: "tps1a ::: 1 = string_to_contents (f1 x)" "transforms M1 (snd (start_config k1 zs)) (T1 n) tps1a" using tps1a_aux zs_def by simp_all lemma length_tps1a [simp]: "length tps1a = k1" using tps1a(2) tm_M1 start_config_length execute_num_tapes transforms_def transits_def turing_machine_def by (smt (verit, del_insts) Suc_1 add_pos_pos less_le_trans less_numeral_extra(1) plus_1_eq_Suc snd_conv) definition tps1b :: "tape list" where "tps1b \ replicate k2 (\[]\, 0)" definition tps1 :: "tape list" where "tps1 \ tps1a @ tps1b" lemma tps1_at_1: "tps1 ! 1 = tps1a ! 1" using tps1_def length_tps1a k_ge by (metis Suc_1 Suc_le_lessD nth_append) lemma tps1_at_1': "tps1 ::: 1 = string_to_contents (f1 x)" using tps1_at_1 tps1a by simp lemma tps1_pos_le: "tps1 :#: 1 \ T1 n" proof - have "execute M1 (start_config k1 zs) (T1 n) = (length M1, tps1a)" using transforms_def transits_def tps1a(2) by (metis (no_types, lifting) execute_after_halting_ge fst_conv start_config_def snd_conv) moreover have "execute M1 (start_config k1 zs) (T1 n) <#> 1 \ T1 n" using head_pos_le_time[OF tm_M1, of 1] k_ge by fastforce ultimately show ?thesis using tps1_at_1 by simp qed lemma length_f1_x: "length (f1 x) \ T1 n" proof - have "execute M1 (start_config k1 zs) (T1 n) = (length M1, tps1a)" using transforms_def transits_def tps1a(2) by (metis (no_types, lifting) execute_after_halting_ge fst_conv start_config_def snd_conv) moreover have "(execute M1 (start_config k1 zs) (T1 n) <:> 1) i = \" if "i > T1 n" for i using blank_after_time[OF that _ _ tm_M1] k_ge(1) by simp ultimately have "(tps1a ::: 1) i = \" if "i > T1 n" for i using that by simp then have "(string_to_contents (f1 x)) i = \" if "i > T1 n" for i using that tps1a(1) by simp then have "length (string_to_symbols (f1 x)) \ T1 n" by (metis length_map order_refl verit_comp_simplify1(3) zero_neq_numeral zero_neq_one) then show ?thesis by simp qed lemma start_config_append: "start_config (k1 + k2) zs = (0, snd (start_config k1 zs) @ tps1b)" proof have "k1 > 0" using tm_M1 turing_machine_def by simp show "fst (start_config (k1 + k2) zs) = fst (0, snd (start_config k1 zs) @ tps1b)" using start_config_def by simp show "snd (start_config (k1 + k2) zs) = snd (0, snd (start_config k1 zs) @ tps1b)" (is "?l = ?r") proof (rule nth_equalityI) have len: "||start_config k1 zs|| = k1" using start_config_length by (simp add: \0 < k1\) show "length ?l = length ?r" using start_config_length tps1b_def tm_M1 turing_machine_def by simp show "?l ! j = ?r ! j" if "j < length ?l" for j proof (cases "j < k1") case True show ?thesis proof (cases "j = 0") case True then show ?thesis using start_config_def `k1 > 0` by simp next case False then have 1: "?l ! j = (\i. if i = 0 then \ else \, 0)" using start_config_def `k1 > 0` True by auto have "?r ! j = snd (start_config k1 zs) ! j" using True len by (simp add: nth_append) then have "?r ! j = (\i. if i = 0 then \ else \, 0)" using start_config4 `k1 > 0` False True by simp then show ?thesis using 1 by simp qed next case False then have j: "j < k1 + k2" "k1 \ j" using that \0 < k1\ add_gr_0 start_config_length by simp_all then have "?r ! j = (\[]\, 0)" using tps1b_def by (simp add: False len nth_append) moreover have "?l ! j = (\i. if i = 0 then \ else \, 0)" using start_config4 `k1 > 0` j by simp ultimately show ?thesis by auto qed qed qed lemma tm1 [transforms_intros]: "transforms tm1 tps0 (T1 n) tps1" proof - let ?M = "append_tapes k1 (k1 + length tps1b) M1" have len: "||start_config k1 zs|| = k1" using start_config_length[of k1 zs] tm_M1 turing_machine_def by simp have "transforms ?M (snd (start_config k1 zs) @ tps1b) (T1 n) (tps1a @ tps1b)" using transforms_append_tapes[OF tm_M1 len tps1a(2), of tps1b] . moreover have "tps0 = snd (start_config k1 zs) @ tps1b" unfolding tps0_def using start_config_append by simp ultimately have *: "transforms ?M tps0 (T1 n) tps1" using tps1_def by simp have "symbols_lt G1 zs" using bit_symbols_zs tm_M1 turing_machine_def by auto moreover have "turing_machine (k1 + k2) G1 ?M" using append_tapes_tm[OF tm_M1, of "k1 + k2"] by (simp add: tps1b_def) ultimately have "transforms (enlarged G1 ?M) tps0 (T1 n) tps1" using transforms_enlarged * tps0_def by simp then show ?thesis using tm1_def tps1b_def by simp qed lemma clean_string_to_contents: "clean_tape (string_to_contents xs, i)" using clean_tape_def by simp definition tps2 :: "tape list" where "tps2 \ tps1 [1 := tps1 ! 1 |#=| 0]" lemma length_tps2 [simp]: "length tps2 = k1 + k2" using tps2_def tps1_def by (simp add: tps1b_def) lemma tm2: assumes "t = Suc (T1 n + tps1 :#: Suc 0)" shows "transforms tm2 tps0 t tps2" unfolding tm2_def proof (tform tps: assms tps2_def) show "1 < length tps1" using tm_M1 turing_machine_def tps1_def by simp show "clean_tape (tps1 ! 1)" using tps1a(1) tps1_at_1 clean_tape_def by simp qed corollary tm2' [transforms_intros]: assumes "t = Suc (2 * T1 n)" shows "transforms tm2 tps0 t tps2" using assms tm2 tps1_pos_le transforms_monotone by simp definition tps3 :: "tape list" where "tps3 \ tps2 [1 := tps2 ! 1 |#=| (Suc (length (f1 x))), k1 := tps2 ! 1 |#=| (Suc (length (f1 x)))]" lemma tm3: assumes "t = Suc (Suc (2 * T1 n + Suc (length (f1 x))))" shows "transforms tm3 tps0 t tps3" unfolding tm3_def proof (tform tps: k_ge) have "Suc 0 < k1 + k2" "0 < k2" using k_ge by simp_all then have *: "tps2 ! 1 = tps1 ! 1 |#=| 0" using tps2_def by (simp add: tps1_def tps1b_def) let ?i = "Suc (length (f1 x))" show "rneigh (tps2 ! 1) {0} ?i" using * tps1_at_1 tps1a by (intro rneighI) auto show "tps3 = tps2 [1 := tps2 ! 1 |+| Suc (length (f1 x)), k1 := implant (tps2 ! 1) (tps2 ! k1) (Suc (length (f1 x)))]" proof - have "tps2 ! 1 |#=| (Suc (length (f1 x))) = tps2 ! 1 |#=| Suc (tps2 :#: 1 + length (f1 x))" by (metis "*" One_nat_def add_Suc plus_1_eq_Suc snd_conv) moreover have "tps2 ! 1 |#=| ?i = implant (tps2 ! 1) (tps2 ! k1) ?i" proof have 1: "tps2 ! 1 = (string_to_contents (f1 x), 0)" using tps1_at_1' * by simp have "tps1 ! k1 = (\[]\, 0)" using tps1_def tps1b_def by (simp add: \0 < k2\ nth_append) then have 2: "tps2 ! k1 = (\[]\, 0)" using tps2_def k_ge by simp then show "snd (tps2 ! 1 |#=| ?i) = snd (implant (tps2 ! 1) (tps2 ! k1) ?i)" using implant by simp have "fst (implant (tps2 ! 1) (tps2 ! k1) ?i) i = fst (tps2 ! 1 |#=| ?i) i" for i using 1 2 implant by simp then show "fst (tps2 ! 1 |#=| ?i) = fst (implant (tps2 ! 1) (tps2 ! k1) ?i)" by auto qed ultimately show ?thesis using tps3_def by simp qed show "t = Suc (2 * T1 n) + Suc (Suc (length (f1 x)))" using assms by simp qed definition "tps3' \ tps1a [1 := (string_to_contents (f1 x), Suc (length (f1 x)))] @ ((string_to_contents (f1 x), Suc (length (f1 x))) # replicate (k2 - 1) (\[]\, 0))" lemma tps3': "tps3 = tps3'" proof (rule nth_equalityI) have "length tps3 = k1 + k2" using tps3_def by simp moreover have "length tps3' = k1 + k2" using k_ge(2) tps3'_def by simp ultimately show "length tps3 = length tps3'" by simp show "tps3 ! j = tps3' ! j" if "j < length tps3" for j proof (cases "j < k1") case True then have rhs: "tps3' ! j = (tps1a [1 := (string_to_contents (f1 x), Suc (length (f1 x)))]) ! j" by (simp add: tps3'_def nth_append) show ?thesis proof (cases "j = 1") case True then have "tps3 ! j = tps2 ! 1 |#=| (Suc (length (f1 x)))" using tps3_def Suc_1 Suc_n_not_le_n \length tps3 = k1 + k2\ k_ge(1) length_tps2 nth_list_update_eq nth_list_update_neq that by auto then have "tps3 ! j = tps1 ! 1 |#=| (Suc (length (f1 x)))" using tps2_def True \length tps3 = k1 + k2\ length_tps2 that by auto then have "tps3 ! j = (string_to_contents (f1 x), Suc (length (f1 x)))" using tps1_at_1 tps1a(1) by simp then show ?thesis using rhs True k_ge(1) by auto next case False then have "tps3 ! j = tps2 ! j" using tps3_def True by simp then have "tps3 ! j = tps1 ! j" using tps2_def False by simp then have "tps3 ! j = tps1a ! j" using length_tps1a tps1_def False True by (simp add: nth_append) moreover have "tps3' ! j = tps1a ! j" using False rhs by simp ultimately show ?thesis by simp qed next case j_ge: False show ?thesis proof (cases "j = k1") case True then have "tps3 ! j = tps2 ! 1 |#=| (Suc (length (f1 x)))" using tps3_def that by simp then have "tps3 ! j = tps1 ! 1 |#=| (Suc (length (f1 x)))" using tps2_def \length tps3 = k1 + k2\ length_tps2 Suc_1 Suc_le_lessD tm1_tm turing_machine_def by simp then have "tps3 ! j = (string_to_contents (f1 x), Suc (length (f1 x)))" using tps1_at_1 tps1a(1) by simp moreover have "tps3' ! j = (string_to_contents (f1 x), Suc (length (f1 x)))" using True tps3'_def by (metis (no_types, lifting) length_list_update length_tps1a nth_append_length) ultimately show ?thesis by simp next case False then have "j > k1" using j_ge by simp then have "tps3' ! j = ((string_to_contents (f1 x), Suc (length (f1 x))) # replicate (k2 - 1) (\[]\, 0)) ! (j - k1)" by (simp add: tps3'_def nth_append) moreover have "j - k1 < k2" by (metis \k1 < j\ \length tps3 = k1 + k2\ add.commute less_diff_conv2 less_imp_le that) ultimately have *: "tps3' ! j = (\[]\, 0)" by (metis (no_types, lifting) Suc_leI \k1 < j\ add_leD1 le_add_diff_inverse2 less_diff_conv2 nth_Cons_pos nth_replicate plus_1_eq_Suc zero_less_diff) have "tps3 ! j = tps2 ! j" using tps3_def \k1 < j\ k_ge(1) by simp then have "tps3 ! j = tps1 ! j" using tps2_def \k1 < j\ k_ge(1) by simp then have "tps3 ! j = tps1b ! (j - k1)" using tps1_def by (simp add: j_ge nth_append) then have "tps3 ! j = (\[]\, 0)" using tps1b_def by (simp add: \j - k1 < k2\) then show ?thesis using * by simp qed qed qed lemma tm3' [transforms_intros]: assumes "t = Suc (Suc (Suc (3 * T1 n)))" shows "transforms tm3 tps0 t tps3'" proof - have "transforms tm3 tps0 (Suc (Suc (2 * T1 n + Suc (length (f1 x))))) tps3" using tm3 by simp moreover have "t \ Suc (Suc (2 * T1 n + Suc (length (f1 x))))" using assms length_f1_x by simp ultimately show ?thesis using tps3' transforms_monotone by auto qed definition "tps4 \ tps1a [1 := (\[]\, 1)] @ ((string_to_contents (f1 x), Suc (length (f1 x))) # replicate (k2 - 1) (\[]\, 0))" lemma tm4: assumes "t = 9 + (3 * T1 n + (Suc (3 * length (string_to_symbols (f1 x)))))" shows "transforms tm4 tps0 t tps4" unfolding tm4_def proof (tform) show "1 < length tps3'" using tps3'_def using tm1_tm turing_machine_def by auto let ?zs = "string_to_symbols (f1 x)" show "proper_symbols ?zs" by simp show "tps4 = tps3'[1 := (\[]\, 1)]" using tps4_def tps3'_def k_ge(1) length_tps1a by (simp add: list_update_append1) show "tps3' ::: 1 = \string_to_symbols (f1 x)\" proof - have "tps3' ! 1 = (string_to_contents (f1 x), Suc (length (f1 x)))" using tps3'_def k_ge(1) length_tps1a by (simp add: nth_append) then show ?thesis by auto qed have "tps3' :#: 1 = Suc (length (f1 x))" using tps3'_def k_ge(1) length_tps1a by (simp add: nth_append) then show "t = Suc (Suc (Suc (3 * T1 n))) + (tps3' :#: 1 + 2 * length (string_to_symbols (f1 x)) + 6)" using assms by simp qed lemma tm4' [transforms_intros]: assumes "t = 10 + (6 * T1 n)" shows "transforms tm4 tps0 t tps4" proof (rule transforms_monotone[OF tm4], simp) show "10 + (3 * T1 n + 3 * length (f1 x)) \ t" using length_f1_x assms by simp qed definition "tps5 \ tps1a [1 := (\[]\, 1)] @ ((string_to_contents (f1 x), 0) # replicate (k2 - 1) (\[]\, 0))" lemma tm5: assumes "t = 11 + (6 * T1 n + tps4 :#: k1)" shows "transforms tm5 tps0 t tps5" unfolding tm5_def proof (tform time: assms) show "k1 < length tps4" using tps4_def length_tps1a by simp show "tps5 = tps4[k1 := tps4 ! k1 |#=| 0]" using tps4_def tps5_def length_tps1a by (metis (no_types, lifting) fst_conv length_list_update list_update_length nth_append_length) show "clean_tape (tps4 ! k1)" using tps4_def length_tps1a clean_tape_def - by (smt (z3) Suc_eq_plus1 add.commute add_cancel_right_right + by (smt (verit) Suc_eq_plus1 add.commute add_cancel_right_right fst_conv length_list_update nat.distinct(1) nat_1_add_1 nth_append_length numeral_3_eq_3) qed abbreviation "ys \ string_to_symbols (f1 x)" abbreviation "m \ length (f1 x)" definition "tps5' \ tps1a [1 := (\[]\, 1)] @ snd (start_config k2 ys)" lemma tps5': "tps5 = tps5'" using tps5_def tps5'_def start_config_def by auto lemma tm5' [transforms_intros]: assumes "t = 12 + 7 * T1 n" shows "transforms tm5 tps0 t tps5'" proof - have "tps4 :#: k1 = Suc (length (f1 x))" using tps4_def by (metis (no_types, lifting) length_list_update length_tps1a nth_append_length snd_conv) then have "tps4 :#: k1 \ Suc (T1 n)" using length_f1_x by simp then have "t \ 11 + (6 * T1 n + tps4 :#: k1)" using assms by simp then show ?thesis using tm5 transforms_monotone tps5' by simp qed definition tps6b :: "tape list" where "tps6b \ SOME tps. tps ::: 1 = string_to_contents (f2 (f1 x)) \ transforms M2 (snd (start_config k2 ys)) (T2 m) tps" lemma tps6b: "tps6b ::: 1 = string_to_contents (f2 (f1 x))" "transforms M2 (snd (start_config k2 ys)) (T2 m) tps6b" using tps6b_def someI_ex[OF computes_in_timeD[OF computes2, of "f1 x"]] by simp_all lemma tps6b_pos_le: "tps6b :#: 1 \ T2 m" proof - have "execute M2 (start_config k2 ys) (T2 m) = (length M2, tps6b)" using transforms_def transits_def tps6b(2) by (metis (no_types, lifting) execute_after_halting_ge fst_conv start_config_def snd_conv) moreover have "execute M2 (start_config k2 ys) (T2 m) <#> 1 \ T2 m" using head_pos_le_time[OF tm_M2, of 1] k_ge by simp ultimately show ?thesis by simp qed lemma length_tps6b: "length tps6b = k2" using tm_M2 execute_num_tapes k_ge(2) tps5' tps5'_def tps5_def tps6b(2) transforms_def transits_def by (smt (verit, ccfv_threshold) One_nat_def Suc_diff_Suc length_Cons length_replicate less_le_trans minus_nat.diff_0 numeral_2_eq_2 prod.sel(2) same_append_eq zero_less_Suc) lemma length_f2_f1_x: "length (f2 (f1 x)) \ T2 m" proof - have "execute M2 (start_config k2 ys) (T2 m) = (length M2, tps6b)" using transforms_def transits_def tps6b(2) by (metis (no_types, lifting) execute_after_halting_ge fst_conv start_config_def snd_conv) moreover have "(execute M2 (start_config k2 ys) (T2 m) <:> 1) i = 0" if "i > T2 m" for i using blank_after_time[OF that _ _ tm_M2] k_ge(2) by simp ultimately have "(tps6b ::: 1) i = \" if "i > T2 m" for i using that by simp then have "(string_to_contents (f2 (f1 x))) i = \" if "i > T2 m" for i using that tps6b(1) by simp then have "length (string_to_symbols (f2 (f1 x))) \ T2 m" by (metis length_map order_refl verit_comp_simplify1(3) zero_neq_numeral zero_neq_one) then show ?thesis by simp qed lemma enlarged_M2: "transforms (enlarged G2 M2) (snd (start_config k2 ys)) (T2 m) tps6b" proof - have "symbols_lt G2 (string_to_symbols (f1 x))" using tm_M2 turing_machine_def by simp then show ?thesis using transforms_enlarged[OF tm_M2 _ tps6b(2)] by simp qed lemma enlarged_M2_tm: "turing_machine k2 G (enlarged G2 M2)" using turing_machine_enlarged tm_M2 G2 by simp definition "tps6 \ tps1a[1 := (\[]\, 1)] @ tps6b" lemma tm56 [transforms_intros]: "transforms tm56 tps5' (T2 m) tps6" using transforms_prepend_tapes[OF enlarged_M2_tm _ _ enlarged_M2, of "tps1a [1 := (\[]\, 1)]" k1] tps5'_def tps6_def tm56_def start_config_length k_ge(2) by auto lemma tps6_at_Suc_k1: "tps6 ::: (k1 + 1) = string_to_contents (f2 (f1 x))" and tps6_pos_le: "tps6 :#: (k1 + 1) \ T2 m" proof - have "tps6 ! (k1 + 1) = tps6b ! 1" using tps6_def length_tps1a length_tps6b by (simp add: nth_append) then show "tps6 ::: (k1 + 1) = string_to_contents (f2 (f1 x))" "tps6 :#: (k1 + 1) \ T2 m" using tps6b(1) tps6b_pos_le by simp_all qed lemma tm6 [transforms_intros]: assumes "t = 12 + 7 * T1 n + T2 m" shows "transforms tm6 tps0 t tps6" unfolding tm6_def by (tform tps: assms) definition "tps7 \ tps1a[1 := (\[]\, 1)] @ tps6b[1 := (string_to_contents (f2 (f1 x)), 1)]" lemma tps7_at_Suc_k1: "tps7 ! (k1 + 1) = (string_to_contents (f2 (f1 x)), 1)" using tps7_def k_ge(2) length_tps1a length_tps6b by (metis (no_types, lifting) One_nat_def Suc_le_lessD add.commute diff_add_inverse length_list_update not_add_less2 nth_append nth_list_update_eq numeral_2_eq_2) lemma tm7: assumes "t = 14 + (7 * T1 n + (T2 m + tps6 :#: Suc k1))" shows "transforms tm7 tps0 t tps7" unfolding tm7_def proof (tform time: assms) show "k1 + 1 < length tps6" using tps6_def k_ge(2) length_tps1a length_tps6b by simp show "clean_tape (tps6 ! (k1 + 1))" using tps6_at_Suc_k1 clean_tape_def by simp show "tps7 = tps6[k1 + 1 := tps6 ! (k1 + 1) |#=| 1]" proof - have "tps6 ! (k1 + 1) |#=| 1 = (string_to_contents (f2 (f1 x)), 1)" using tps6_at_Suc_k1 by simp then show ?thesis using tps6_def tps7_def length_tps1a length_tps6b k_ge tps7_at_Suc_k1 by (metis (no_types, lifting) add.commute diff_add_inverse length_list_update list_update_append not_add_less2 plus_1_eq_Suc) qed qed corollary tm7' [transforms_intros]: assumes "t = 14 + 7 * T1 n + 2 * T2 m" shows "transforms tm7 tps0 t tps7" proof (rule transforms_monotone[OF tm7], simp) show "14 + (7 * T1 n + (T2 (length (f1 x)) + tps6 :#: Suc k1)) \ t" using assms tps6_pos_le by simp qed definition "tps8 \ tps1a[1 := (string_to_contents (f2 (f1 x)), Suc (length (f2 (f1 x))))] @ tps6b[1 := (string_to_contents (f2 (f1 x)), Suc (length (f2 (f1 x))))]" lemma tps8_at_1: "tps8 ::: 1 = string_to_contents (f2 (f1 x))" using tps8_def length_tps1a k_ge(1) by (metis (no_types, lifting) One_nat_def Suc_le_lessD length_list_update nth_append nth_list_update_eq numeral_2_eq_2 prod.sel(1)) lemma tm8: assumes "t = 15 + 7 * T1 n + 2 * T2 m + length (f2 (f1 x))" shows "transforms tm8 tps0 t tps8" unfolding tm8_def proof (tform tps: assms) show "k1 + 1 < length tps7" using tps7_def length_tps1a length_tps6b k_ge(2) by simp show "1 < length tps7" using tps7_def length_tps6b k_ge(2) by simp let ?n = "length (f2 (f1 x))" show "rneigh (tps7 ! (k1 + 1)) {\} ?n" proof (rule rneighI) show "(tps7 ::: (k1 + 1)) (tps7 :#: (k1 + 1) + ?n) \ {\}" using tps7_at_Suc_k1 by simp show "\n'. n' < ?n \ (tps7 ::: (k1 + 1)) (tps7 :#: (k1 + 1) + n') \ {\}" using tps7_at_Suc_k1 by simp qed show "tps8 = tps7 [k1 + 1 := tps7 ! (k1 + 1) |+| ?n, 1 := implant (tps7 ! (k1 + 1)) (tps7 ! 1) ?n]" (is "tps8 = ?tps") proof (rule nth_equalityI) show "length tps8 = length ?tps" using tps8_def tps7_def by simp have len: "length tps8 = k1 + k2" using tps8_def length_tps6b by simp show "tps8 ! j = ?tps ! j" if "j < length tps8" for j proof (cases "j < k1") case j_less: True then have lhs: "tps8 ! j = tps1a[1 := (string_to_contents (f2 (f1 x)), Suc (length (f2 (f1 x))))] ! j" using tps8_def length_tps1a length_tps6b k_ge by (simp add: nth_append) show ?thesis proof (cases "j = 1") case True then have 1: "?tps ! j = implant (tps7 ! (k1 + 1)) (tps7 ! 1) ?n" using \1 < length tps7\ by simp have 2: "tps8 ! j = (string_to_contents (f2 (f1 x)), Suc (length (f2 (f1 x))))" using lhs True j_less by simp have 3: "tps7 ! 1 = (\[]\, 1)" using tps7_def length_tps1a by (metis (no_types, lifting) True j_less length_list_update nth_append nth_list_update_eq) have "implant (string_to_contents (f2 (f1 x)), 1) (\[]\, 1) ?n = (string_to_contents (f2 (f1 x)), Suc ?n)" using implant contents_def by auto then show ?thesis using 1 2 3 tps7_at_Suc_k1 by simp next case False then have "?tps ! j = tps7 ! j" by (metis One_nat_def Suc_eq_plus1 add.commute j_less not_add_less2 nth_list_update_neq) then have "?tps ! j = tps1a ! j" using False j_less tps7_def length_tps1a by (metis (no_types, lifting) length_list_update nth_append nth_list_update_neq) moreover have "tps8 ! j = tps1a ! j" using False j_less tps8_def lhs by simp ultimately show ?thesis by simp qed next case j_ge: False then have lhs: "tps8 ! j = tps6b[1 := (string_to_contents (f2 (f1 x)), Suc (length (f2 (f1 x))))] ! (j - k1)" using tps8_def length_tps1a length_tps6b k_ge by (simp add: nth_append) show ?thesis proof (cases "j = Suc k1") case True then have "tps8 ! j = (string_to_contents (f2 (f1 x)), Suc (length (f2 (f1 x))))" using lhs len that length_tps6b by simp moreover have "?tps ! j = tps7 ! Suc k1 |+| ?n" using True \k1 + 1 < length tps7\ k_ge(1) by simp ultimately show ?thesis using tps7_at_Suc_k1 True by simp next case False then have "tps8 ! j = tps6b ! (j - k1)" using lhs by simp moreover have "?tps ! j = tps7 ! j" using False j_ge that k_ge(1) by simp ultimately show ?thesis using tps7_def j_ge False length_tps1a by (metis (no_types, lifting) add.commute add_diff_inverse_nat length_list_update nth_append nth_list_update_neq plus_1_eq_Suc) qed qed qed qed lemma tm8': assumes "t = 15 + 7 * T1 n + 3 * T2 m" shows "transforms tm8 tps0 t tps8" proof (rule transforms_monotone[OF tm8], simp) show "15 + 7 * T1 n + 2 * T2 m + length (f2 (f1 x)) \ t" using length_f2_f1_x assms by simp qed lemma tm8'_mono: assumes "mono T2" and "t = 15 + 7 * T1 n + 3 * T2 (T1 n)" shows "transforms tm8 tps0 t tps8" proof (rule transforms_monotone[OF tm8'], simp) have "T2 (T1 n) \ T2 m" using assms(1) length_f1_x monoE by auto then show "15 + 7 * T1 n + 3 * T2 m \ t" using assms(2) by simp qed end (* context x *) lemma computes_composed_mono: assumes "mono T2" and "T = (\n. 15 + 7 * T1 n + 3 * T2 (T1 n))" shows "computes_in_time (k1 + k2) tm8 (f2 \ f1) T" proof fix x have "tps8 x ::: 1 = string_to_contents (f2 (f1 x))" using tps8_at_1 by simp moreover have "transforms tm8 (snd (start_config (k1 + k2) (string_to_symbols x))) (T (length x)) (tps8 x)" using tm8'_mono assms tps0_def zs_def by simp ultimately show "\tps. tps ::: 1 = string_to_contents ((f2 \ f1) x) \ transforms tm8 (snd (start_config (k1 + k2) (string_to_symbols x))) (T (length x)) tps" by force qed end (* locale compose *) lemma computes_composed_poly: assumes tm_M1: "turing_machine k1 G1 M1" and tm_M2: "turing_machine k2 G2 M2" and computes1: "computes_in_time k1 M1 f1 T1" and computes2: "computes_in_time k2 M2 f2 T2" assumes "big_oh_poly T1" and "big_oh_poly T2" shows "\T k G M. big_oh_poly T \ turing_machine k G M \ computes_in_time k M (f2 \ f1) T" proof - obtain d1 :: nat where "big_oh T1 (\n. n ^ d1)" using assms(5) big_oh_poly_def by auto obtain b c d2 :: nat where cm: "d2 > 0" "\n. T2 n \ b + c * n ^ d2" using big_oh_poly_offset[OF assms(6)] by auto let ?U = "\n. b + c * n ^ d2" have U: "T2 n \ ?U n" for n using cm by simp have "mono ?U" by standard (simp add: cm(1)) have computesU: "computes_in_time k2 M2 f2 ?U" using computes_in_time_mono[OF computes2 U] . interpret compo: compose k1 k2 G1 G2 M1 M2 T1 ?U f1 f2 using assms computesU compose.intro by simp let ?M = "compo.tm8" let ?T = "(\n. 15 + 7 * T1 n + 3 * (b + c * T1 n ^ d2))" have "computes_in_time (k1 + k2) ?M (f2 \ f1) ?T" using compo.computes_composed_mono[OF `mono ?U`, of ?T] by simp moreover have "big_oh_poly ?T" proof - have "big_oh_poly (\n. n ^ d2)" using big_oh_poly_poly by simp moreover have "(\n. T1 n ^ d2) = (\n. n ^ d2) \ T1" by auto ultimately have "big_oh_poly (\n. T1 n ^ d2)" using big_oh_poly_composition[OF assms(5)] by auto then have "big_oh_poly (\n. 3 * (b + c * T1 n ^ d2))" using big_oh_poly_const big_oh_poly_prod big_oh_poly_sum by simp then show ?thesis using assms(5) big_oh_poly_const big_oh_poly_prod big_oh_poly_sum by simp qed moreover have "turing_machine (k1 + k2) compo.G ?M" using compo.tm8_tm . ultimately show ?thesis by auto qed end diff --git a/thys/Cook_Levin/Elementary.thy b/thys/Cook_Levin/Elementary.thy --- a/thys/Cook_Levin/Elementary.thy +++ b/thys/Cook_Levin/Elementary.thy @@ -1,2416 +1,2416 @@ section \Elementary Turing machines\label{s:tm-elementary}\ theory Elementary imports Combinations begin text \ In this section we devise some simple yet useful Turing machines. We have already fully analyzed the empty TM, where start and halting state coincide, in the lemmas~@{thm [source] computes_Nil_empty}, @{thm [source] Nil_tm}, and @{thm [source] transforms_Nil}. The next more complex TMs are those with exactly one command. They represent TMs with two states: the halting state and the start state where the action happens. This action might last for one step only, or the TM may stay in this state for longer; for example, it can move a tape head rightward to the next blank symbol. We will also start using the @{text ";;"} operator to combine some of the one-command TMs. \ text \ Most Turing machines we are going to construct throughout this section and indeed the entire article are really families of Turing machines that usually are parameterized by tape indices. \ type_synonym tapeidx = nat text \ Throughout this article, names of commands are prefixed with @{text cmd_} and names of Turing machines with @{text tm_}. Usually for a TM named @{term tm_foo} there is a lemma @{text tm_foo_tm} stating that it really is a Turing machine and a lemma @{text transforms_tm_fooI} describing its semantics and running time. The lemma usually receives a @{text transforms_intros} attribute for use with our proof method. If @{term "tm_foo"} comprises more than two TMs we will typically analyze the semantics and running time in a locale named @{text "turing_machine_foo"}. The first example of this is @{term tm_equals} in Section~\ref{s:tm-elementary-comparing}. When it comes to running times, we will have almost no scruples simplifying upper bounds to have the form $a + b\cdot n^c$ for some constants $a, b, c$, even if this means, for example, bounding $n \log n$ by $n^2$. \ subsection \Clean tapes\ text \ Most of our Turing machines will not change the start symbol in the first cell of a tape nor will they write the start symbol anywhere else. The only exceptions are machines that simulate arbitrary other machines. We call tapes that have the start symbol only in the first cell \emph{clean tapes}. \ definition clean_tape :: "tape \ bool" where "clean_tape tp \ \i. fst tp i = \ \ i = 0" lemma clean_tapeI: assumes "\i. fst tp i = \ \ i = 0" shows "clean_tape tp" unfolding clean_tape_def using assms by simp lemma clean_tapeI': assumes "fst tp 0 = \" and "\i. i > 0 \ fst tp i \ \" shows "clean_tape tp" unfolding clean_tape_def using assms by auto text \ A clean configuration is one with only clean tapes. \ definition clean_config :: "config \ bool" where "clean_config cfg \ (\j<||cfg||. \i. (cfg <:> j) i = \ \ i = 0)" lemma clean_config_tapes: "clean_config cfg = (\tp\set (snd cfg). clean_tape tp)" using clean_config_def clean_tape_def by (metis in_set_conv_nth) lemma clean_configI: assumes "\j i. j < length tps \ fst (tps ! j) i = \ \ i = 0" shows "clean_config (q, tps)" unfolding clean_config_def using assms by simp lemma clean_configI': assumes "\tp i. tp \ set tps \ fst tp i = \ \ i = 0" shows "clean_config (q, tps)" using clean_configI assms by simp lemma clean_configI'': assumes "\tp. tp \ set tps \ (fst tp 0 = \ \ (\i>0. fst tp i \ \))" shows "clean_config (q, tps)" using clean_configI' assms by blast lemma clean_configE: assumes "clean_config (q, tps)" shows "\j i. j < length tps \ fst (tps ! j) i = \ \ i = 0" using clean_config_def assms by simp lemma clean_configE': assumes "clean_config (q, tps)" shows "\tp i. tp \ set tps \ fst tp i = \ \ i = 0" using clean_configE assms by (metis in_set_conv_nth) lemma clean_contents_proper [simp]: "proper_symbols zs \ clean_tape (\zs\, q)" using clean_tape_def contents_def proper_symbols_ne1 by simp lemma contents_clean_tape': "proper_symbols zs \ fst tp = \zs\ \ clean_tape tp" using contents_def clean_tape_def by (simp add: nat_neq_iff) text \ Some more lemmas about @{const contents}: \ lemma contents_append_blanks: "\ys @ replicate m \\ = \ys\" proof fix i consider "i = 0" | "0 < i \ i \ length ys" | "length ys < i \ i \ length ys + m" | "length ys + m < i" by linarith then show "\ys @ replicate m \\ i = \ys\ i" proof (cases) case 1 then show ?thesis by simp next case 2 then show ?thesis using contents_inbounds by (metis (no_types, opaque_lifting) Suc_diff_1 Suc_le_eq le_add1 le_trans length_append nth_append) next case 3 then show ?thesis - by (smt (z3) Suc_diff_Suc add_diff_inverse_nat contents_def diff_Suc_1 diff_commute leD less_one + by (smt (verit) Suc_diff_Suc add_diff_inverse_nat contents_def diff_Suc_1 diff_commute leD less_one less_or_eq_imp_le nat_add_left_cancel_le not_less_eq nth_append nth_replicate) next case 4 then show ?thesis by simp qed qed lemma contents_append_update: assumes "length ys = m" shows "\ys @ [v] @ zs\(Suc m := w) = \ys @ [w] @ zs\" proof fix i consider "i = 0" | "0 < i \ i < Suc m" | "i = Suc m" | "i > Suc m \ i \ Suc m + length zs" | "i > Suc m + length zs" by linarith then show "(\ys @ [v] @ zs\(Suc m := w)) i = \ys @ [w] @ zs\ i" (is "?l = ?r") proof (cases) case 1 then show ?thesis by simp next case 2 then have "?l = (ys @ [v] @ zs) ! (i - 1)" using assms contents_inbounds by simp then have *: "?l = ys ! (i - 1)" using 2 assms by (metis Suc_diff_1 Suc_le_lessD less_Suc_eq_le nth_append) have "?r = (ys @ [w] @ zs) ! (i - 1)" using 2 assms contents_inbounds by simp then have "?r = ys ! (i - 1)" using 2 assms by (metis Suc_diff_1 Suc_le_lessD less_Suc_eq_le nth_append) then show ?thesis using * by simp next case 3 then show ?thesis using assms by auto next case 4 then have "?l = (ys @ [v] @ zs) ! (i - 1)" using assms contents_inbounds by simp then have "?l = ((ys @ [v]) @ zs) ! (i - 1)" by simp then have *: "?l = zs ! (i - 1 - Suc m)" using 4 assms by (metis diff_Suc_1 length_append_singleton less_imp_Suc_add not_add_less1 nth_append) then have "?r = (ys @ [w] @ zs) ! (i - 1)" using 4 assms contents_inbounds by simp then have "?r = ((ys @ [w]) @ zs) ! (i - 1)" by simp then have "?r = zs ! (i - 1 - Suc m)" using 4 assms by (metis diff_Suc_1 length_append_singleton less_imp_Suc_add not_add_less1 nth_append) then show ?thesis using * by simp next case 5 then show ?thesis using assms by simp qed qed lemma contents_snoc: "\ys\(Suc (length ys) := w) = \ys @ [w]\" proof fix i consider "i = 0" | "0 < i \ i \ length ys" | "i = Suc (length ys)" | "i > Suc (length ys)" by linarith then show "(\ys\(Suc (length ys) := w)) i = \ys @ [w]\ i" proof (cases) case 1 then show ?thesis by simp next case 2 then show ?thesis by (smt (verit, ccfv_SIG) contents_def diff_Suc_1 ex_least_nat_less fun_upd_apply leD le_Suc_eq length_append_singleton less_imp_Suc_add nth_append) next case 3 then show ?thesis by simp next case 4 then show ?thesis by simp qed qed definition config_update_pos :: "config \ nat \ nat \ config" where "config_update_pos cfg j p \ (fst cfg, (snd cfg)[j:=(cfg <:> j, p)])" lemma config_update_pos_0: "config_update_pos cfg j (cfg <#> j) = cfg" using config_update_pos_def by simp definition config_update_fwd :: "config \ nat \ nat \ config" where "config_update_fwd cfg j d \ (fst cfg, (snd cfg)[j:=(cfg <:> j, cfg <#> j + d)])" lemma config_update_fwd_0: "config_update_fwd cfg j 0 = cfg" using config_update_fwd_def by simp lemma config_update_fwd_additive: "config_update_fwd (config_update_fwd cfg j d1) j d2 = (config_update_fwd cfg j (d1 + d2))" using config_update_fwd_def - by (smt add.commute add.left_commute fst_conv le_less_linear list_update_beyond list_update_overwrite nth_list_update_eq sndI) + by (smt (verit) add.commute add.left_commute fst_conv le_less_linear list_update_beyond list_update_overwrite nth_list_update_eq sndI) subsection \Moving tape heads\ text \ Among the most simple things a Turing machine can do is moving one of its tape heads. \ subsubsection \Moving left\ text \ The next command makes a TM move its head on tape $j$ one cell to the left unless, of course, it is in the leftmost cell already. \ definition cmd_left :: "tapeidx \ command" where "cmd_left j \ \rs. (1, map (\i. (rs ! i, if i = j then Left else Stay)) [0.. (cmd_left j rs) [!] j = (rs ! j, Left)" using cmd_left_def by simp lemma cmd_left''': "i < length rs \ i \ j \ (cmd_left j rs) [!] i = (rs ! i, Stay)" using cmd_left_def by simp lemma tape_list_eq: assumes "length tps' = length tps" and "\i. i < length tps \ i \ j \ tps' ! i = tps ! i" and "tps' ! j = x" shows "tps' = tps[j := x]" - using assms by (smt length_list_update list_update_beyond not_le nth_equalityI nth_list_update) + using assms by (smt (verit) length_list_update list_update_beyond not_le nth_equalityI nth_list_update) lemma sem_cmd_left: assumes "j < length tps" shows "sem (cmd_left j) (0, tps) = (1, tps[j:=(fst (tps ! j), snd (tps ! j) - 1)])" proof show "fst (sem (cmd_left j) (0, tps)) = fst (1, tps[j := (fst (tps ! j), snd (tps ! j) - 1)])" using cmd_left' sem_fst by simp have "snd (sem (cmd_left j) (0, tps)) = tps[j := (fst (tps ! j), snd (tps ! j) - 1)]" proof (rule tape_list_eq) show "||sem (cmd_left j) (0, tps)|| = length tps" using turing_command_left sem_num_tapes2' by (metis snd_eqD) show "sem (cmd_left j) (0, tps) i = tps ! i" if "i < length tps" and "i \ j" for i proof - let ?rs = "read tps" have "length ?rs = length tps" using read_length by simp moreover have "sem (cmd_left j) (0, tps) i = act (cmd_left j ?rs [!] i) (tps ! i)" by (simp add: cmd_left_def sem_snd that(1)) ultimately show ?thesis using that act_Stay cmd_left''' by simp qed show "sem (cmd_left j) (0, tps) j = (fst (tps ! j), snd (tps ! j) - 1)" using assms act_Left cmd_left_def read_length sem_snd by simp qed then show "snd (sem (cmd_left j) (0, tps)) = snd (1, tps[j := (fst (tps ! j), snd (tps ! j) - 1)])" by simp qed definition tm_left :: "tapeidx \ machine" where "tm_left j \ [cmd_left j]" lemma tm_left_tm: "k \ 2 \ G \ 4 \ turing_machine k G (tm_left j)" unfolding tm_left_def using turing_command_left by auto lemma exe_tm_left: assumes "j < length tps" shows "exe (tm_left j) (0, tps) = (1, tps[j := tps ! j |-| 1])" unfolding tm_left_def using sem_cmd_left assms by (simp add: exe_lt_length) lemma execute_tm_left: assumes "j < length tps" shows "execute (tm_left j) (0, tps) (Suc 0) = (1, tps[j := tps ! j |-| 1])" using assms exe_tm_left by simp lemma transits_tm_left: assumes "j < length tps" shows "transits (tm_left j) (0, tps) (Suc 0) (1, tps[j := tps ! j |-| 1])" using execute_tm_left assms transitsI by blast lemma transforms_tm_left: assumes "j < length tps" shows "transforms (tm_left j) tps (Suc 0) (tps[j := tps ! j |-| 1])" using transits_tm_left assms by (simp add: tm_left_def transforms_def) lemma transforms_tm_leftI [transforms_intros]: assumes "j < length tps" and "t = 1" and "tps' = tps[j := tps ! j |-| 1]" shows "transforms (tm_left j) tps t tps'" using transits_tm_left assms by (simp add: tm_left_def transforms_def) subsubsection \Moving right\ text \ The next command makes the head on tape $j$ move one cell to the right. \ definition cmd_right :: "tapeidx \ command" where "cmd_right j \ \rs. (1, map (\i. (rs ! i, if i = j then Right else Stay)) [0.. (cmd_right j rs) [!] j = (rs ! j, Right)" using cmd_right_def by simp lemma cmd_right''': "i < length rs \ i \ j \ (cmd_right j rs) [!] i = (rs ! i, Stay)" using cmd_right_def by simp lemma sem_cmd_right: assumes "j < length tps" shows "sem (cmd_right j) (0, tps) = (1, tps[j:=(fst (tps ! j), snd (tps ! j) + 1)])" proof show "fst (sem (cmd_right j) (0, tps)) = fst (1, tps[j := (fst (tps ! j), snd (tps ! j) + 1)])" using cmd_right' sem_fst by simp have "snd (sem (cmd_right j) (0, tps)) = tps[j := (fst (tps ! j), snd (tps ! j) + 1)]" proof (rule tape_list_eq) show "||sem (cmd_right j) (0, tps)|| = length tps" using sem_num_tapes3 turing_command_right by (metis snd_conv) show "sem (cmd_right j) (0, tps) i = tps ! i" if "i < length tps" and "i \ j" for i proof - let ?rs = "read tps" have "length ?rs = length tps" using read_length by simp moreover have "sem (cmd_right j) (0, tps) i = act (cmd_right j ?rs [!] i) (tps ! i)" by (simp add: cmd_right_def sem_snd that(1)) ultimately show ?thesis using that act_Stay cmd_right''' by simp qed show "sem (cmd_right j) (0, tps) j = (fst (tps ! j), snd (tps ! j) + 1)" using assms act_Right cmd_right_def read_length sem_snd by simp qed then show "snd (sem (cmd_right j) (0, tps)) = snd (1, tps[j := (fst (tps ! j), snd (tps ! j) + 1)])" by simp qed definition tm_right :: "tapeidx \ machine" where "tm_right j \ [cmd_right j]" lemma tm_right_tm: "k \ 2 \ G \ 4 \ turing_machine k G (tm_right j)" unfolding tm_right_def using turing_command_right cmd_right' by auto lemma exe_tm_right: assumes "j < length tps" shows "exe (tm_right j) (0, tps) = (1, tps[j:=(fst (tps ! j), snd (tps ! j) + 1)])" unfolding tm_right_def using sem_cmd_right assms by (simp add: exe_lt_length) lemma execute_tm_right: assumes "j < length tps" shows "execute (tm_right j) (0, tps) (Suc 0) = (1, tps[j:=(fst (tps ! j), snd (tps ! j) + 1)])" using assms exe_tm_right by simp lemma transits_tm_right: assumes "j < length tps" shows "transits (tm_right j) (0, tps) (Suc 0) (1, tps[j:=(fst (tps ! j), snd (tps ! j) + 1)])" using execute_tm_right assms transitsI by blast lemma transforms_tm_right: assumes "j < length tps" shows "transforms (tm_right j) tps (Suc 0) (tps[j := tps ! j |+| 1])" using transits_tm_right assms by (simp add: tm_right_def transforms_def) lemma transforms_tm_rightI [transforms_intros]: assumes "j < length tps" and "t = Suc 0" and "tps' = tps[j := tps ! j |+| 1]" shows "transforms (tm_right j) tps t tps'" using transits_tm_right assms by (simp add: tm_right_def transforms_def) subsubsection \Moving right on several tapes\ text \ The next command makes the heads on all tapes from a set $J$ of tapes move one cell to the right. \ definition cmd_right_many :: "tapeidx set \ command" where "cmd_right_many J \ \rs. (1, map (\i. (rs ! i, if i \ J then Right else Stay)) [0..j. if j \ J then tps ! j |+| 1 else tps ! j) [0..j. if j \ J then tps ! j |+| 1 else tps ! j) [0..j. if j \ J then tps ! j |+| 1 else tps ! j) [0.. J") case True let ?rs = "read tps" have "length ?rs = length tps" using read_length by simp moreover have "sem (cmd_right_many J) (0, tps) j = act (cmd_right_many J ?rs [!] j) (tps ! j)" using cmd_right_many_def sem_snd that True len by auto moreover have "?rhs ! j = tps ! j |+| 1" using that len True by simp ultimately show ?thesis using that act_Right cmd_right_many_def True len by simp next case False let ?rs = "read tps" have "length ?rs = length tps" using read_length by simp moreover have "sem (cmd_right_many J) (0, tps) j = act (cmd_right_many J ?rs [!] j) (tps ! j)" using cmd_right_many_def sem_snd that False len by auto moreover have "?rhs ! j = tps ! j" using that len False by simp ultimately show ?thesis using that act_Stay cmd_right_many_def False len by simp qed qed then show "snd (sem (cmd_right_many J) (0, tps)) = snd (1, map (\j. if j \ J then tps ! j |+| 1 else tps ! j) [0.. machine" where "tm_right_many J \ [cmd_right_many J]" lemma tm_right_many_tm: "k \ 2 \ G \ 4 \ turing_machine k G (tm_right_many J)" unfolding tm_right_many_def using turing_command_right_many by auto lemma transforms_tm_right_manyI [transforms_intros]: assumes "t = Suc 0" and "tps' = map (\j. if j \ J then tps ! j |+| 1 else tps ! j) [0..j. if j \ J then tps ! j |+| 1 else tps ! j) [0..j. if j \ J then tps ! j |+| 1 else tps ! j) [0..j. if j \ J then tps ! j |+| 1 else tps ! j) [0..j. if j \ J then tps ! j |+| 1 else tps ! j) [0..Copying and translating tape contents\ text \ The Turing machines in this section scan a tape $j_1$ and copy the symbols to another tape $j_2$. Scanning can be performed in either direction, and ``copying'' may include mapping the symbols. \ subsubsection \Copying and translating from one tape to another\ text \ The next predicate is true iff.\ on the given tape the next symbol from the set $H$ of symbols is exactly $n$ cells to the right from the current head position. Thus, a command that moves the tape head right until it finds a symbol from $H$ takes $n$ steps and moves the head $n$ cells right. \ definition rneigh :: "tape \ symbol set \ nat \ bool" where "rneigh tp H n \ fst tp (snd tp + n) \ H \ (\n' < n. fst tp (snd tp + n') \ H)" lemma rneighI: assumes "fst tp (snd tp + n) \ H" and "\n'. n' < n \ fst tp (snd tp + n') \ H" shows "rneigh tp H n" using assms rneigh_def by simp text \ The analogous predicate for moving to the left: \ definition lneigh :: "tape \ symbol set \ nat \ bool" where "lneigh tp H n \ fst tp (snd tp - n) \ H \ (\n' < n. fst tp (snd tp - n') \ H)" lemma lneighI: assumes "fst tp (snd tp - n) \ H" and "\n'. n' < n \ fst tp (snd tp - n') \ H" shows "lneigh tp H n" using assms lneigh_def by simp text \ The next command scans tape $j_1$ rightward until it reaches a symbol from the set $H$. While doing so it copies the symbols, after applying a mapping $f$, to tape $j_2$. \ definition cmd_trans_until :: "tapeidx \ tapeidx \ symbol set \ (symbol \ symbol) \ command" where "cmd_trans_until j1 j2 H f \ \rs. if rs ! j1 \ H then (1, map (\r. (r, Stay)) rs) else (0, map (\i. (if i = j2 then f (rs ! j1) else rs ! i, if i = j1 \ i = j2 then Right else Stay)) [0.. The analogous command for moving to the left: \ definition cmd_ltrans_until :: "tapeidx \ tapeidx \ symbol set \ (symbol \ symbol) \ command" where "cmd_ltrans_until j1 j2 H f \ \rs. if rs ! j1 \ H then (1, map (\r. (r, Stay)) rs) else (0, map (\i. (if i = j2 then f (rs ! j1) else rs ! i, if i = j1 \ i = j2 then Left else Stay)) [0.. j1 \ H" shows "sem (cmd_trans_until j1 j2 H f) (0, tps) = (1, tps)" using cmd_trans_until_def tapes_at_read read_length assms act_Stay by (intro semI[OF proper_cmd_trans_until]) auto lemma sem_cmd_ltrans_until_1: assumes "j1 < k" and "length tps = k" and "(0, tps) <.> j1 \ H" shows "sem (cmd_ltrans_until j1 j2 H f) (0, tps) = (1, tps)" using cmd_ltrans_until_def tapes_at_read read_length assms act_Stay by (intro semI[OF proper_cmd_ltrans_until]) auto lemma sem_cmd_trans_until_2: assumes "j1 < k" and "length tps = k" and "(0, tps) <.> j1 \ H" shows "sem (cmd_trans_until j1 j2 H f) (0, tps) = (0, tps[j1 := tps ! j1 |+| 1, j2 := tps ! j2 |:=| (f (tps :.: j1)) |+| 1])" using cmd_trans_until_def tapes_at_read read_length assms act_Stay act_Right by (intro semI[OF proper_cmd_trans_until]) auto lemma sem_cmd_ltrans_until_2: assumes "j1 < k" and "length tps = k" and "(0, tps) <.> j1 \ H" shows "sem (cmd_ltrans_until j1 j2 H f) (0, tps) = (0, tps[j1 := tps ! j1 |-| 1, j2 := tps ! j2 |:=| (f (tps :.: j1)) |-| 1])" using cmd_ltrans_until_def tapes_at_read read_length assms act_Stay act_Left by (intro semI[OF proper_cmd_ltrans_until]) auto definition tm_trans_until :: "tapeidx \ tapeidx \ symbol set \ (symbol \ symbol) \ machine" where "tm_trans_until j1 j2 H f \ [cmd_trans_until j1 j2 H f]" definition tm_ltrans_until :: "tapeidx \ tapeidx \ symbol set \ (symbol \ symbol) \ machine" where "tm_ltrans_until j1 j2 H f \ [cmd_ltrans_until j1 j2 H f]" lemma tm_trans_until_tm: assumes "0 < j2" and "j1 < k" and "j2 < k" and "\h 2" and "G \ 4" shows "turing_machine k G (tm_trans_until j1 j2 H f)" unfolding tm_trans_until_def cmd_trans_until_def using assms turing_machine_def by auto lemma tm_ltrans_until_tm: assumes "0 < j2" and "j1 < k" and "j2 < k" and "\h 2" and "G \ 4" shows "turing_machine k G (tm_ltrans_until j1 j2 H f)" unfolding tm_ltrans_until_def cmd_ltrans_until_def using assms turing_machine_def by auto lemma exe_tm_trans_until_1: assumes "j1 < k" and "length tps = k" and "(0, tps) <.> j1 \ H" shows "exe (tm_trans_until j1 j2 H f) (0, tps) = (1, tps)" unfolding tm_trans_until_def using sem_cmd_trans_until_1 assms exe_lt_length by simp lemma exe_tm_ltrans_until_1: assumes "j1 < k" and "length tps = k" and "(0, tps) <.> j1 \ H" shows "exe (tm_ltrans_until j1 j2 H f) (0, tps) = (1, tps)" unfolding tm_ltrans_until_def using sem_cmd_ltrans_until_1 assms exe_lt_length by simp lemma exe_tm_trans_until_2: assumes "j1 < k" and "length tps = k" and "(0, tps) <.> j1 \ H" shows "exe (tm_trans_until j1 j2 H f) (0, tps) = (0, tps[j1 := tps ! j1 |+| 1, j2 := tps ! j2 |:=| (f (tps :.: j1)) |+| 1])" unfolding tm_trans_until_def using sem_cmd_trans_until_2 assms exe_lt_length by simp lemma exe_tm_ltrans_until_2: assumes "j1 < k" and "length tps = k" and "(0, tps) <.> j1 \ H" shows "exe (tm_ltrans_until j1 j2 H f) (0, tps) = (0, tps[j1 := tps ! j1 |-| 1, j2 := tps ! j2 |:=| (f (tps :.: j1)) |-| 1])" unfolding tm_ltrans_until_def using sem_cmd_ltrans_until_2 assms exe_lt_length by simp text \ Let $tp_1$ and $tp_2$ be two tapes with head positions $i_1$ and $i_2$, respectively. The next function describes the result of overwriting the symbols at positions $i_2, \dots, i_2 + n - 1$ on tape $tp_2$ by the symbols at positions $i_1, \dots, i_1 + n - 1$ on tape $tp_1$ after applying a symbol map $f$. \ definition transplant :: "tape \ tape \ (symbol \ symbol)\ nat \ tape" where "transplant tp1 tp2 f n \ (\i. if snd tp2 \ i \ i < snd tp2 + n then f (fst tp1 (snd tp1 + i - snd tp2)) else fst tp2 i, snd tp2 + n)" text \ The analogous function for moving to the left while copying: \ definition ltransplant :: "tape \ tape \ (symbol \ symbol)\ nat \ tape" where "ltransplant tp1 tp2 f n \ (\i. if snd tp2 - n < i \ i \ snd tp2 then f (fst tp1 (snd tp1 + i - snd tp2)) else fst tp2 i, snd tp2 - n)" lemma transplant_0: "transplant tp1 tp2 f 0 = tp2" unfolding transplant_def by standard auto lemma ltransplant_0: "ltransplant tp1 tp2 f 0 = tp2" unfolding ltransplant_def by standard auto lemma transplant_upd: "transplant tp1 tp2 f n |:=| (f ( |.| (tp1 |+| n))) |+| 1 = transplant tp1 tp2 f (Suc n)" unfolding transplant_def by auto lemma ltransplant_upd: assumes "n < snd tp2" shows "ltransplant tp1 tp2 f n |:=| (f ( |.| (tp1 |-| n))) |-| 1 = ltransplant tp1 tp2 f (Suc n)" unfolding ltransplant_def using assms by fastforce lemma tapes_ltransplant_upd: assumes "t < tps :#: j2" and "t < tps :#: j1" and "j1 < k" and "j2 < k" and "length tps = k" and "tps' = tps[j1 := tps ! j1 |-| t, j2 := ltransplant (tps ! j1) (tps ! j2) f t]" shows "tps'[j1 := tps' ! j1 |-| 1, j2 := tps' ! j2 |:=| (f (tps' :.: j1)) |-| 1] = tps[j1 := tps ! j1 |-| Suc t, j2 := ltransplant (tps ! j1) (tps ! j2) f (Suc t)]" (is "?lhs = ?rhs") proof (rule nth_equalityI) show 1: "length ?lhs = length ?rhs" using assms by simp have len: "length ?lhs = k" using assms by simp show "?lhs ! j = ?rhs ! j" if "j < length ?lhs" for j proof - have "j < k" using len that by simp show ?thesis proof (cases "j \ j1 \ j \ j2") case True then show ?thesis using assms by simp next case j1j2: False show ?thesis proof (cases "j1 = j2") case True then have j: "j = j1" "j = j2" using j1j2 by simp_all have "tps' ! j1 = ltransplant (tps ! j1) (tps ! j2) f t" using j assms that by simp then have *: "snd (tps' ! j1) = snd (tps ! j1) - t" using j ltransplant_def by simp then have "fst (tps' ! j1) = (\i. if snd (tps ! j2) - t < i \ i \ snd (tps ! j2) then f (fst (tps ! j1) (snd (tps ! j1) + i - snd (tps ! j2))) else fst (tps ! j2) i)" using j ltransplant_def assms by auto then have "fst (tps' ! j1) = (\i. if snd (tps ! j1) - t < i \ i \ snd (tps ! j1) then f (fst (tps ! j1) (snd (tps ! j1) + i - snd (tps ! j1))) else fst (tps ! j1) i)" using j by auto then have "fst (tps' ! j1) (snd (tps ! j1) - t) = fst (tps ! j1) (snd (tps ! j1) - t)" by simp then have "tps' :.: j1 = fst (tps ! j1) (snd (tps ! j1) - t)" using * by simp then have "?lhs ! j = (ltransplant (tps ! j1) (tps ! j2) f t) |:=| (f ( |.| (tps ! j1 |-| t))) |-| 1" using assms(6) j that by simp then have "?lhs ! j = (ltransplant (tps ! j1) (tps ! j2) f (Suc t))" using ltransplant_upd assms(1) by simp moreover have "?rhs ! j = ltransplant (tps ! j1) (tps ! j2) f (Suc t)" using assms(6) that j by simp ultimately show ?thesis by simp next case j1_neq_j2: False then show ?thesis proof (cases "j = j1") case True then have "?lhs ! j = tps' ! j1 |-| 1" using assms j1_neq_j2 by simp then have "?lhs ! j = (tps ! j1 |-| t) |-| 1" using assms j1_neq_j2 by simp moreover have "?rhs ! j = tps ! j1 |-| Suc t" using True assms j1_neq_j2 by simp ultimately show ?thesis by simp next case False then have j: "j = j2" using j1j2 by simp then have "?lhs ! j = tps' ! j2 |:=| (f (tps' :.: j1)) |-| 1" using assms by simp then have "?lhs ! j = (ltransplant (tps ! j1) (tps ! j2) f t) |:=| (f (tps' :.: j1)) |-| 1" using assms by simp then have "?lhs ! j = (ltransplant (tps ! j1) (tps ! j2) f (Suc t))" - using ltransplant_def assms ltransplant_upd by (smt (z3) j1_neq_j2 nth_list_update_eq nth_list_update_neq) + using ltransplant_def assms ltransplant_upd by (smt (verit) j1_neq_j2 nth_list_update_eq nth_list_update_neq) moreover have "?rhs ! j = ltransplant (tps ! j1) (tps ! j2) f (Suc t)" using assms(6) that j by simp ultimately show ?thesis by simp qed qed qed qed qed lemma execute_tm_trans_until_less: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" and "t \ n" shows "execute (tm_trans_until j1 j2 H f) (0, tps) t = (0, tps[j1 := tps ! j1 |+| t, j2 := transplant (tps ! j1) (tps ! j2) f t])" using assms(5) proof (induction t) case 0 then show ?case using transplant_0 by simp next case (Suc t) then have "t < n" by simp let ?M = "tm_trans_until j1 j2 H f" have "execute ?M (0, tps) (Suc t) = exe ?M (execute ?M (0, tps) t)" by simp also have "... = exe ?M (0, tps[j1 := tps ! j1 |+| t, j2 := transplant (tps ! j1) (tps ! j2) f t])" (is "_ = exe ?M (0, ?tps)") using Suc by simp also have "... = (0, ?tps[j1 := ?tps ! j1 |+| 1, j2 := ?tps ! j2 |:=| (f (?tps :.: j1)) |+| 1])" proof (rule exe_tm_trans_until_2[where ?k=k]) show "j1 < k" using assms(1) . show "length (tps[j1 := tps ! j1 |+| t, j2 := transplant (tps ! j1) (tps ! j2) f t]) = k" using assms by simp show "(0, tps[j1 := tps ! j1 |+| t, j2 := transplant (tps ! j1) (tps ! j2) f t]) <.> j1 \ H" using assms transplant_def rneigh_def \t < n\ - by (smt fst_conv length_list_update less_not_refl2 nth_list_update_eq nth_list_update_neq snd_conv) + by (smt (verit) fst_conv length_list_update less_not_refl2 nth_list_update_eq nth_list_update_neq snd_conv) qed finally show ?case using assms transplant_upd by auto - (smt add.commute fst_conv transplant_def transplant_upd less_not_refl2 list_update_overwrite list_update_swap + (smt (verit) add.commute fst_conv transplant_def transplant_upd less_not_refl2 list_update_overwrite list_update_swap nth_list_update_eq nth_list_update_neq plus_1_eq_Suc snd_conv) qed lemma execute_tm_ltrans_until_less: assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "t \ n" and "n \ tps :#: j1" and "n \ tps :#: j2" shows "execute (tm_ltrans_until j1 j2 H f) (0, tps) t = (0, tps[j1 := tps ! j1 |-| t, j2 := ltransplant (tps ! j1) (tps ! j2) f t])" using assms(5) proof (induction t) case 0 then show ?case using ltransplant_0 by simp next case (Suc t) then have "t < n" by simp have 1: "t < tps :#: j2" using assms(7) Suc by simp have 2: "t < tps :#: j1" using assms(6) Suc by simp let ?M = "tm_ltrans_until j1 j2 H f" define tps' where "tps' = tps[j1 := tps ! j1 |-| t, j2 := ltransplant (tps ! j1) (tps ! j2) f t]" have "execute ?M (0, tps) (Suc t) = exe ?M (execute ?M (0, tps) t)" by simp also have "... = exe ?M (0, tps')" using Suc tps'_def by simp also have "... = (0, tps'[j1 := tps' ! j1 |-| 1, j2 := tps' ! j2 |:=| (f (tps' :.: j1)) |-| 1])" proof (rule exe_tm_ltrans_until_2[where ?k=k]) show "j1 < k" using assms(1) . show "length tps' = k" using assms tps'_def by simp show "(0, tps') <.> j1 \ H" using assms ltransplant_def tps'_def lneigh_def \t < n\ - by (smt fst_conv length_list_update less_not_refl2 nth_list_update_eq nth_list_update_neq snd_conv) + by (smt (verit) fst_conv length_list_update less_not_refl2 nth_list_update_eq nth_list_update_neq snd_conv) qed finally show ?case using tapes_ltransplant_upd[OF 1 2 assms(1,2,3) tps'_def] by simp qed lemma execute_tm_trans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" shows "execute (tm_trans_until j1 j2 H f) (0, tps) (Suc n) = (1, tps[j1 := tps ! j1 |+| n, j2 := transplant (tps ! j1) (tps ! j2) f n])" proof - let ?M = "tm_trans_until j1 j2 H f" have "execute ?M (0, tps) (Suc n) = exe ?M (execute ?M (0, tps) n)" by simp also have "... = exe ?M (0, tps[j1 := tps ! j1 |+| n, j2 := transplant (tps ! j1) (tps ! j2) f n])" using execute_tm_trans_until_less[where ?t=n] assms by simp also have "... = (1, tps[j1 := tps ! j1 |+| n, j2 := transplant (tps ! j1) (tps ! j2) f n])" (is "_ = (1, ?tps)") proof - have "length ?tps = k" using assms(3) by simp moreover have "(0, ?tps) <.> j1 \ H" using rneigh_def transplant_def assms - by (smt fst_conv length_list_update less_not_refl2 nth_list_update_eq nth_list_update_neq snd_conv) + by (smt (verit) fst_conv length_list_update less_not_refl2 nth_list_update_eq nth_list_update_neq snd_conv) ultimately show ?thesis using exe_tm_trans_until_1 assms by simp qed finally show ?thesis by simp qed lemma execute_tm_ltrans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "n \ tps :#: j1" and "n \ tps :#: j2" shows "execute (tm_ltrans_until j1 j2 H f) (0, tps) (Suc n) = (1, tps[j1 := tps ! j1 |-| n, j2 := ltransplant (tps ! j1) (tps ! j2) f n])" proof - let ?M = "tm_ltrans_until j1 j2 H f" have "execute ?M (0, tps) (Suc n) = exe ?M (execute ?M (0, tps) n)" by simp also have "... = exe ?M (0, tps[j1 := tps ! j1 |-| n, j2 := ltransplant (tps ! j1) (tps ! j2) f n])" using execute_tm_ltrans_until_less[where ?t=n] assms by simp also have "... = (1, tps[j1 := tps ! j1 |-| n, j2 := ltransplant (tps ! j1) (tps ! j2) f n])" (is "_ = (1, ?tps)") proof - have "length ?tps = k" using assms(3) by simp moreover have "(0, ?tps) <.> j1 \ H" using lneigh_def ltransplant_def assms by (smt (verit, ccfv_threshold) fst_conv length_list_update less_not_refl nth_list_update_eq nth_list_update_neq snd_conv) ultimately show ?thesis using exe_tm_ltrans_until_1 assms by simp qed finally show ?thesis by simp qed lemma transits_tm_trans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" shows "transits (tm_trans_until j1 j2 H f) (0, tps) (Suc n) (1, tps[j1 := tps ! j1 |+| n, j2 := transplant (tps ! j1) (tps ! j2) f n])" using execute_tm_trans_until[OF assms] transitsI[of _ _ "Suc n" _ "Suc n"] by blast lemma transits_tm_ltrans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "n \ tps :#: j1" and "n \ tps :#: j2" shows "transits (tm_ltrans_until j1 j2 H f) (0, tps) (Suc n) (1, tps[j1 := tps ! j1 |-| n, j2 := ltransplant (tps ! j1) (tps ! j2) f n])" using execute_tm_ltrans_until[OF assms] transitsI[of _ _ "Suc n" _ "Suc n"] by blast lemma transforms_tm_trans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" shows "transforms (tm_trans_until j1 j2 H f) tps (Suc n) (tps[j1 := tps ! j1 |+| n, j2 := transplant (tps ! j1) (tps ! j2) f n])" using tm_trans_until_def transforms_def transits_tm_trans_until[OF assms] by simp lemma transforms_tm_ltrans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "n \ tps :#: j1" and "n \ tps :#: j2" shows "transforms (tm_ltrans_until j1 j2 H f) tps (Suc n) (tps[j1 := tps ! j1 |-| n, j2 := ltransplant (tps ! j1) (tps ! j2) f n])" using tm_ltrans_until_def transforms_def transits_tm_ltrans_until[OF assms] by simp corollary transforms_tm_trans_untilI [transforms_intros]: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" and "t = Suc n" and "tps' = tps[j1 := tps ! j1 |+| n, j2 := transplant (tps ! j1) (tps ! j2) f n]" shows "transforms (tm_trans_until j1 j2 H f) tps t tps'" using transforms_tm_trans_until[OF assms(1-4)] assms(5,6) by simp corollary transforms_tm_ltrans_untilI [transforms_intros]: assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "n \ tps :#: j1" and "n \ tps :#: j2" and "t = Suc n" and "tps' = tps[j1 := tps ! j1 |-| n, j2 := ltransplant (tps ! j1) (tps ! j2) f n]" shows "transforms (tm_ltrans_until j1 j2 H f) tps t tps'" using transforms_tm_ltrans_until[OF assms(1-6)] assms(7,8) by simp subsubsection \Copying one tape to another\ text \ If we set the symbol map $f$ in @{const tm_trans_until} to the identity function, we get a Turing machine that simply makes a copy. \ definition tm_cp_until :: "tapeidx \ tapeidx \ symbol set \ machine" where "tm_cp_until j1 j2 H \ tm_trans_until j1 j2 H id" lemma id_symbol: "\h symbol) h < G" by simp lemma tm_cp_until_tm: assumes "0 < j2" and "j1 < k" and "j2 < k" and "G \ 4" shows "turing_machine k G (tm_cp_until j1 j2 H)" unfolding tm_cp_until_def using tm_trans_until_tm id_symbol assms turing_machine_def by simp abbreviation implant :: "tape \ tape \ nat \ tape" where "implant tp1 tp2 n \ transplant tp1 tp2 id n" lemma implant: "implant tp1 tp2 n = (\i. if snd tp2 \ i \ i < snd tp2 + n then fst tp1 (snd tp1 + i - snd tp2) else fst tp2 i, snd tp2 + n)" using transplant_def by auto lemma implantI: assumes "tp' = (\i. if snd tp2 \ i \ i < snd tp2 + n then fst tp1 (snd tp1 + i - snd tp2) else fst tp2 i, snd tp2 + n)" shows "implant tp1 tp2 n = tp'" using implant assms by simp lemma fst_snd_pair: "fst t = a \ snd t = b \ t = (a, b)" by auto lemma implantI': assumes "fst tp' = (\i. if snd tp2 \ i \ i < snd tp2 + n then fst tp1 (snd tp1 + i - snd tp2) else fst tp2 i)" and "snd tp' = snd tp2 + n" shows "implant tp1 tp2 n = tp'" using implantI fst_snd_pair[OF assms] by simp lemma implantI'': assumes "\i. snd tp2 \ i \ i < snd tp2 + n \ fst tp' i = fst tp1 (snd tp1 + i - snd tp2)" and "\i. i < snd tp2 \ fst tp' i = fst tp2 i" and "\i. snd tp2 + n \ i \ fst tp' i = fst tp2 i" assumes "snd tp' = snd tp2 + n" shows "implant tp1 tp2 n = tp'" using assms implantI' by (meson linorder_le_less_linear) lemma implantI''': assumes "\i. i2 \ i \ i < i2 + n \ ys i = ys1 (i1 + i - i2)" and "\i. i < i2 \ ys i = ys2 i" and "\i. i2 + n \ i \ ys i = ys2 i" assumes "i = i2 + n" shows "implant (ys1, i1) (ys2, i2) n = (ys, i)" using assms implantI'' by auto lemma implant_self: "implant tp tp n = tp |+| n" unfolding transplant_def by auto lemma transforms_tm_cp_untilI [transforms_intros]: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" and "t = Suc n" and "tps' = tps[j1 := tps ! j1 |+| n, j2 := implant (tps ! j1) (tps ! j2) n]" shows "transforms (tm_cp_until j1 j2 H) tps t tps'" unfolding tm_cp_until_def using transforms_tm_trans_untilI[OF assms(1-6)] by simp lemma implant_contents: assumes "i > 0" and "n + (i - 1) \ length xs" shows "implant (\xs\, i) (\ys\, Suc (length ys)) n = (\ys @ (take n (drop (i - 1) xs))\, Suc (length ys) + n)" (is "?lhs = ?rhs") proof - have lhs: "?lhs = (\j. if Suc (length ys) \ j \ j < Suc (length ys) + n then \xs\ (i + j - Suc (length ys)) else \ys\ j, Suc (length ys) + n)" using implant by auto let ?zs = "ys @ (take n (drop (i - 1) xs))" have lenzs: "length ?zs = length ys + n" using assms by simp have fst_rhs: "fst ?rhs = (\j. if j = 0 then 1 else if j \ length ys + n then ?zs ! (j - 1) else 0)" using assms by auto have "(\j. if Suc (length ys) \ j \ j < Suc (length ys) + n then \xs\ (i + j - Suc (length ys)) else \ys\ j) = (\j. if j = 0 then 1 else if j \ length ys + n then ?zs ! (j - 1) else 0)" (is "?l = ?r") proof fix j consider "j = 0" | "j > 0 \ j \ length ys" | "j > length ys \ j \ length ys + n" | "j > length ys + n" by linarith then show "?l j = ?r j" proof (cases) case 1 then show ?thesis by simp next case 2 then show ?thesis - using assms contents_def by (smt (z3) Suc_diff_1 less_trans_Suc not_add_less1 not_le not_less_eq_eq nth_append) + using assms contents_def by (smt (verit) Suc_diff_1 less_trans_Suc not_add_less1 not_le not_less_eq_eq nth_append) next case 3 then have "?r j = ?zs ! (j - 1)" by simp also have "... = take n (drop (i - 1) xs) ! (j - 1 - length ys)" using 3 lenzs by (metis add.right_neutral diff_is_0_eq le_add_diff_inverse not_add_less2 not_le not_less_eq nth_append plus_1_eq_Suc) also have "... = take n (drop (i - 1) xs) ! (j - Suc (length ys))" by simp also have "... = xs ! (i - 1 + j - Suc (length ys))" using 3 assms by auto also have "... = \xs\ (i + j - Suc (length ys))" using assms contents_def 3 by auto also have "... = ?l j" using 3 by simp finally have "?r j = ?l j" . then show ?thesis by simp next case 4 then show ?thesis by simp qed qed then show ?thesis using lhs fst_rhs by simp qed subsubsection \Moving to the next specific symbol\ text \ Copying a tape to itself means just moving to the right. \ definition tm_right_until :: "tapeidx \ symbol set \ machine" where "tm_right_until j H \ tm_cp_until j j H" text \ Copying a tape to itself does not change the tape. So this is a Turing machine even for the input tape $j = 0$, unlike @{const tm_cp_until} where the target tape cannot, in general, be the input tape. \ lemma tm_right_until_tm: assumes "j < k" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_right_until j H)" unfolding tm_right_until_def tm_cp_until_def tm_trans_until_def cmd_trans_until_def using assms turing_machine_def by auto lemma transforms_tm_right_untilI [transforms_intros]: assumes "j < length tps" and "rneigh (tps ! j) H n" and "t = Suc n" and "tps' = (tps[j := tps ! j |+| n])" shows "transforms (tm_right_until j H) tps t tps'" using transforms_tm_cp_untilI assms implant_self tm_right_until_def by (metis list_update_id nth_list_update_eq) subsubsection \Translating to a constant symbol\ text \ Another way to specialize @{const tm_trans_until} and @{const tm_ltrans_until} is to have a constant function $f$. \ definition tm_const_until :: "tapeidx \ tapeidx \ symbol set \ symbol \ machine" where "tm_const_until j1 j2 H h \ tm_trans_until j1 j2 H (\_. h)" lemma tm_const_until_tm: assumes "0 < j2" and "j1 < k" and "j2 < k" and "h < G" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_const_until j1 j2 H h)" unfolding tm_const_until_def using tm_trans_until_tm assms turing_machine_def by metis text \ Continuing with our fantasy names ending in \emph{-plant}, we name the operation @{term constplant}. \ abbreviation constplant :: "tape \ symbol \ nat \ tape" where "constplant tp2 h n \ transplant (\_. 0, 0) tp2 (\_. h) n" lemma constplant_transplant: "constplant tp2 h n = transplant tp1 tp2 (\_. h) n" using transplant_def by simp lemma constplant: "constplant tp2 h n = (\i. if snd tp2 \ i \ i < snd tp2 + n then h else fst tp2 i, snd tp2 + n)" using transplant_def by simp lemma transforms_tm_const_untilI [transforms_intros]: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" and "t = Suc n" and "tps' = tps[j1 := tps ! j1 |+| n, j2 := constplant (tps ! j2) h n]" shows "transforms (tm_const_until j1 j2 H h) tps t tps'" unfolding tm_const_until_def using transforms_tm_trans_untilI assms constplant_transplant by metis definition tm_lconst_until :: "tapeidx \ tapeidx \ symbol set \ symbol \ machine" where "tm_lconst_until j1 j2 H h \ tm_ltrans_until j1 j2 H (\_. h)" lemma tm_lconst_until_tm: assumes "0 < j2" and "j1 < k" and "j2 < k" and "h < G" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_lconst_until j1 j2 H h)" unfolding tm_lconst_until_def using tm_ltrans_until_tm assms turing_machine_def by metis abbreviation lconstplant :: "tape \ symbol \ nat \ tape" where "lconstplant tp2 h n \ ltransplant (\_. 0, 0) tp2 (\_. h) n" lemma lconstplant_ltransplant: "lconstplant tp2 h n = ltransplant tp1 tp2 (\_. h) n" using ltransplant_def by simp lemma lconstplant: "lconstplant tp2 h n = (\i. if snd tp2 - n < i \ i \ snd tp2 then h else fst tp2 i, snd tp2 - n)" using ltransplant_def by simp lemma transforms_tm_lconst_untilI [transforms_intros]: assumes "0 < j2" and "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "n \ tps :#: j1" and "n \ tps :#: j2" and "t = Suc n" and "tps' = tps[j1 := tps ! j1 |-| n, j2 := lconstplant (tps ! j2) h n]" shows "transforms (tm_lconst_until j1 j2 H h) tps t tps'" unfolding tm_lconst_until_def using transforms_tm_ltrans_untilI assms lconstplant_ltransplant by metis subsection \Writing single symbols\ text \ The next command writes a fixed symbol $h$ to tape $j$. It does not move a tape head. \ definition cmd_write :: "tapeidx \ symbol \ command" where "cmd_write j h rs \ (1, map (\i. (if i = j then h else rs ! i, Stay)) [0.. symbol \ machine" where "tm_write j h \ [cmd_write j h]" lemma tm_write_tm: assumes "0 < j" and "j < k" and "h < G" and "G \ 4" shows "turing_machine k G (tm_write j h)" unfolding tm_write_def cmd_write_def using assms turing_machine_def by auto lemma transforms_tm_writeI [transforms_intros]: assumes "tps' = tps[j := tps ! j |:=| h]" shows "transforms (tm_write j h) tps 1 tps'" unfolding tm_write_def using sem_cmd_write exe_lt_length assms tm_write_def transits_def transforms_def by auto text \ The next command writes the symbol to tape $j_2$ that results from applying a function $f$ to the symbol read from tape $j_1$. It does not move any tape heads. \ definition cmd_trans2 :: "tapeidx \ tapeidx \ (symbol \ symbol) \ command" where "cmd_trans2 j1 j2 f rs \ (1, map (\i. (if i = j2 then f (rs ! j1) else rs ! i, Stay)) [0.. tapeidx \ (symbol \ symbol) \ machine" where "tm_trans2 j1 j2 f \ [cmd_trans2 j1 j2 f]" lemma tm_trans2_tm: assumes "j1 < k" and "0 < j2" and "j2 < k" and "\h < G. f h < G" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_trans2 j1 j2 f)" unfolding tm_trans2_def cmd_trans2_def using assms turing_machine_def by auto lemma exe_tm_trans2: assumes "j1 < length tps" shows "exe (tm_trans2 j1 j2 f) (0, tps) = (1, tps[j2 := tps ! j2 |:=| (f (tps :.: j1))])" unfolding tm_trans2_def using sem_cmd_trans2 assms exe_lt_length by simp lemma execute_tm_trans2: assumes "j1 < length tps" shows "execute (tm_trans2 j1 j2 f) (0, tps) 1 = (1, tps[j2 := tps ! j2 |:=| (f (tps :.: j1))])" using assms exe_tm_trans2 by simp lemma transits_tm_trans2: assumes "j1 < length tps" shows "transits (tm_trans2 j1 j2 f) (0, tps) 1 (1, tps[j2 := tps ! j2 |:=| (f (tps :.: j1))])" using assms execute_tm_trans2 transits_def by auto lemma transforms_tm_trans2: assumes "j1 < length tps" shows "transforms (tm_trans2 j1 j2 f) tps 1 (tps[j2 := tps ! j2 |:=| (f (tps :.: j1))])" using tm_trans2_def assms transits_tm_trans2 transforms_def by simp lemma transforms_tm_trans2I [transforms_intros]: assumes "j1 < length tps" and "tps' = tps[j2 := tps ! j2 |:=| (f (tps :.: j1))]" shows "transforms (tm_trans2 j1 j2 f) tps 1 tps'" using assms transforms_tm_trans2 by simp text \ Equating the two tapes in @{const tm_trans2}, we can map a symbol in-place. \ definition tm_trans :: "tapeidx \ (symbol \ symbol) \ machine" where "tm_trans j f \ tm_trans2 j j f" lemma tm_trans_tm: assumes "0 < j" and "j < k" and "\h < G. f h < G" and "G \ 4" shows "turing_machine k G (tm_trans j f)" unfolding tm_trans_def using tm_trans2_tm assms by simp lemma transforms_tm_transI [transforms_intros]: assumes "j < length tps" and "tps' = tps[j := tps ! j |:=| (f (tps :.: j))]" shows "transforms (tm_trans j f) tps 1 tps'" using assms transforms_tm_trans2 tm_trans_def by simp text \ The next command is like the previous one, except it also moves the tape head to the right. \ definition cmd_rtrans :: "tapeidx \ (symbol \ symbol) \ command" where "cmd_rtrans j f rs \ (1, map (\i. (if i = j then f (rs ! i) else rs ! i, if i = j then Right else Stay)) [0.. (symbol \ symbol) \ machine" where "tm_rtrans j f \ [cmd_rtrans j f]" lemma tm_rtrans_tm: assumes "0 < j" and "j < k" and "\h 2" and "G \ 4" shows "turing_machine k G (tm_rtrans j f)" unfolding tm_rtrans_def cmd_rtrans_def using assms turing_machine_def by auto lemma exe_tm_rtrans: assumes "j < length tps" shows "exe (tm_rtrans j f) (0, tps) = (1, tps[j := tps ! j |:=| (f (tps :.: j)) |+| 1])" unfolding tm_rtrans_def using sem_cmd_rtrans assms exe_lt_length by simp lemma execute_tm_rtrans: assumes "j < length tps" shows "execute (tm_rtrans j f) (0, tps) 1 = (1, tps[j := tps ! j |:=| (f (tps :.: j)) |+| 1])" using assms exe_tm_rtrans by simp lemma transits_tm_rtrans: assumes "j < length tps" shows "transits (tm_rtrans j f) (0, tps) 1 (1, tps[j := tps ! j |:=| (f (tps :.: j)) |+| 1])" using assms execute_tm_rtrans transits_def by auto lemma transforms_tm_rtrans: assumes "j < length tps" shows "transforms (tm_rtrans j f) tps 1 (tps[j := tps ! j |:=| (f (tps :.: j)) |+| 1])" using assms transits_tm_rtrans transforms_def by (metis One_nat_def length_Cons list.size(3) tm_rtrans_def) lemma transforms_tm_rtransI [transforms_intros]: assumes "j < length tps" and "tps' = tps[j := tps ! j |:=| (f (tps :.: j)) |+| 1]" shows "transforms (tm_rtrans j f) tps 1 tps'" using assms transforms_tm_rtrans by simp text \ The next command writes a fixed symbol $h$ to all tapes in the set $J$. \ definition cmd_write_many :: "tapeidx set \ symbol \ command" where "cmd_write_many J h rs \ (1, map (\i. (if i \ J then h else rs ! i, Stay)) [0..j. if j \ J then tps ! j |:=| h else tps ! j) [0.. symbol \ machine" where "tm_write_many J h \ [cmd_write_many J h]" lemma tm_write_many_tm: assumes "0 \ J" and "\j\J. j < k" and "h < G" and "k \ 2" and "G \ 4" shows "turing_machine k G (tm_write_many J h)" unfolding tm_write_many_def cmd_write_many_def using assms turing_machine_def by auto lemma exe_tm_write_many: "exe (tm_write_many J h) (0, tps) = (1, map (\j. if j \ J then tps ! j |:=| h else tps ! j) [0..j. if j \ J then tps ! j |:=| h else tps ! j) [0..j. if j \ J then tps ! j |:=| h else tps ! j) [0..j\J. j < k" and "length tps = k" and "length tps' = k" and "\j. j \ J \ tps' ! j = tps ! j |:=| h" and "\j. j < k \ j \ J \ tps' ! j = tps ! j" shows "transforms (tm_write_many J h) tps 1 tps'" proof - have "tps' = map (\j. if j \ J then tps ! j |:=| h else tps ! j) [0..Writing a symbol multiple times\ text \ In this section we devise a Turing machine that writes the symbol sequence $h^m$ with a hard-coded symbol $h$ and number $m$ to a tape. The resulting tape is defined by the next operation: \ definition overwrite :: "tape \ symbol \ nat \ tape" where "overwrite tp h m \ (\i. if snd tp \ i \ i < snd tp + m then h else fst tp i, snd tp + m)" lemma overwrite_0: "overwrite tp h 0 = tp" proof - have "snd (overwrite tp h 0) = snd tp" unfolding overwrite_def by simp moreover have "fst (overwrite tp h 0) = fst tp" unfolding overwrite_def by auto ultimately show ?thesis using prod_eqI by blast qed lemma overwrite_upd: "(overwrite tp h t) |:=| h |+| 1 = overwrite tp h (Suc t)" using overwrite_def by auto lemma overwrite_upd': assumes "j < length tps" and "tps' = tps[j := overwrite (tps ! j) h t]" shows "(tps[j := overwrite (tps ! j) h t])[j := tps' ! j |:=| h |+| 1] = tps[j := overwrite (tps ! j) h (Suc t)]" using assms overwrite_upd by simp text \ The next command writes the symbol $h$ to the tape $j$ and moves the tape head to the right. \ definition cmd_char :: "tapeidx \ symbol \ command" where "cmd_char j z = cmd_rtrans j (\_. z)" lemma turing_command_char: assumes "0 < j" and "j < k" and "h < G" shows "turing_command k 1 G (cmd_char j h)" unfolding cmd_char_def cmd_rtrans_def using assms by auto definition tm_char :: "tapeidx \ symbol \ machine" where "tm_char j z \ [cmd_char j z]" lemma tm_char_tm: assumes "0 < j" and "j < k" and "G \ 4" and "z < G" shows "turing_machine k G (tm_char j z)" using assms turing_command_char tm_char_def by auto lemma transforms_tm_charI [transforms_intros]: assumes "j < length tps" and "tps' = tps[j := tps ! j |:=| z |+| 1]" shows "transforms (tm_char j z) tps 1 tps'" using assms transforms_tm_rtransI tm_char_def cmd_char_def tm_rtrans_def by metis lemma sem_cmd_char: assumes "j < length tps" shows "sem (cmd_char j h) (0, tps) = (1, tps[j := tps ! j |:=| h |+| 1])" using cmd_char_def cmd_rtrans_def tapes_at_read read_length assms act_Right by (intro semI) auto text \ The next TM is a sequence of $m$ @{const cmd_char} commands properly relocated. It writes a sequence of $m$ times the symbol $h$ to tape $j$. \ definition tm_write_repeat :: "tapeidx \ symbol \ nat \ machine" where "tm_write_repeat j h m \ map (\i. relocate_cmd i (cmd_char j h)) [0.. 2" and "G \ 4" shows "turing_machine k G (tm_write_repeat j h m)" proof let ?M = "tm_write_repeat j h m" show "2 \ k" and "4 \ G" using assms(4,5) . show "turing_command k (length ?M) G (?M ! i)" if "i < length ?M" for i proof - have "?M ! i = relocate_cmd i (cmd_char j h)" using that tm_write_repeat_def by simp then have "turing_command k (1 + i) G (?M ! i)" using assms turing_command_char turing_command_relocate_cmd by metis then show ?thesis using turing_command_mono that by simp qed qed lemma exe_tm_write_repeat: assumes "j < length tps" and "q < m" shows "exe (tm_write_repeat j h m) (q, tps) = (Suc q, tps[j := tps ! j |:=| h |+| 1])" using sem_cmd_char assms sem sem_relocate_cmd tm_write_repeat_def by (auto simp add: exe_lt_length) lemma execute_tm_write_repeat: assumes "j < length tps" and "t \ m" shows "execute (tm_write_repeat j h m) (0, tps) t = (t, tps[j := overwrite (tps ! j) h t])" using assms(2) proof (induction t) case 0 then show ?case using overwrite_0 by simp next case (Suc t) then have "t < m" by simp have "execute (tm_write_repeat j h m) (0, tps) (Suc t) = exe (tm_write_repeat j h m) (execute (tm_write_repeat j h m) (0, tps) t)" by simp also have "... = exe (tm_write_repeat j h m) (t, tps[j := overwrite (tps ! j) h t])" using Suc by simp also have "... = (Suc t, tps[j := overwrite (tps ! j) h (Suc t)])" using `t < m` exe_tm_write_repeat assms overwrite_upd' by simp finally show ?case . qed lemma transforms_tm_write_repeatI [transforms_intros]: assumes "j < length tps" and "tps' = tps[j := overwrite (tps ! j) h m]" shows "transforms (tm_write_repeat j h m) tps m tps'" using assms execute_tm_write_repeat transits_def transforms_def tm_write_repeat_def by auto subsection \Moving to the start of the tape\ text \ The next command moves the head on tape $j$ to the left until it reaches a symbol from the set $H$: \ definition cmd_left_until :: "symbol set \ tapeidx \ command" where "cmd_left_until H j rs \ if rs ! j \ H then (1, map (\r. (r, Stay)) rs) else (0, map (\i. (rs ! i, if i = j then Left else Stay)) [0.. j \ H" shows "sem (cmd_left_until H j) (0, tps) = (1, tps)" using cmd_left_until_def tapes_at_read read_length assms act_Stay by (intro semI) auto lemma sem_cmd_left_until_2: assumes "j < k" and "length tps = k" and "(0, tps) <.> j \ H" shows "sem (cmd_left_until H j) (0, tps) = (0, tps[j := tps ! j |-| 1])" using cmd_left_until_def tapes_at_read read_length assms act_Stay act_Left by (intro semI) auto definition tm_left_until :: "symbol set \ tapeidx \ machine" where "tm_left_until H j \ [cmd_left_until H j]" lemma tm_left_until_tm: assumes "k \ 2" and "G \ 4" shows "turing_machine k G (tm_left_until H j)" unfolding tm_left_until_def cmd_left_until_def using assms turing_machine_def by auto text \ A \emph{begin tape} for a set of symbols has one of these symbols only in cell zero. It generalizes the concept of clean tapes, where the set of symbols is $\{\triangleright\}$. \ definition begin_tape :: "symbol set \ tape \ bool" where "begin_tape H tp \ \i. fst tp i \ H \ i = 0" lemma begin_tapeI: assumes "fst tp 0 \ H" and "\i. i > 0 \ fst tp i \ H" shows "begin_tape H tp" unfolding begin_tape_def using assms by auto lemma exe_tm_left_until_1: assumes "j < length tps" and "(0, tps) <.> j \ H" shows "exe (tm_left_until H j) (0, tps) = (1, tps)" using tm_left_until_def assms exe_lt_length sem_cmd_left_until_1 by auto lemma exe_tm_left_until_2: assumes "j < length tps" and "(0, tps) <.> j \ H" shows "exe (tm_left_until H j) (0, tps) = (0, tps[j := tps ! j |-| 1])" using tm_left_until_def assms exe_lt_length sem_cmd_left_until_2 by auto text \ We do not show the semantics of @{const tm_left_until} for the general case, but only for when applied to begin tapes. \ lemma execute_tm_left_until_less: assumes "j < length tps" and "begin_tape H (tps ! j)" and "t \ tps :#: j" shows "execute (tm_left_until H j) (0, tps) t = (0, tps[j := tps ! j |-| t])" using assms(3) proof (induction t) case 0 then show ?case by simp next case (Suc t) then have "fst (tps ! j) (snd (tps ! j) - t) \ H" using assms begin_tape_def by simp then have neq_0: "fst (tps ! j |-| t) (snd (tps ! j |-| t)) \ H" by simp have "execute (tm_left_until H j) (0, tps) (Suc t) = exe (tm_left_until H j) (execute (tm_left_until H j) (0, tps) t)" by simp also have "... = exe (tm_left_until H j) (0, tps[j := tps ! j |-| t])" using Suc by simp also have "... = (0, tps[j := tps ! j |-| (Suc t)])" using neq_0 exe_tm_left_until_2 assms by simp finally show ?case by simp qed lemma execute_tm_left_until: assumes "j < length tps" and "begin_tape H (tps ! j)" shows "execute (tm_left_until H j) (0, tps) (Suc (tps :#: j)) = (1, tps[j := tps ! j |#=| 0])" using assms begin_tape_def exe_tm_left_until_1 execute_tm_left_until_less by simp lemma transits_tm_left_until: assumes "j < length tps" and "begin_tape H (tps ! j)" shows "transits (tm_left_until H j) (0, tps) (Suc (tps :#: j)) (1, tps[j := tps ! j |#=| 0])" using execute_imp_transits[OF execute_tm_left_until[OF assms]] by simp lemma transforms_tm_left_until: assumes "j < length tps" and "begin_tape H (tps ! j)" shows "transforms (tm_left_until H j) tps (Suc (tps :#: j)) (tps[j := tps ! j |#=| 0])" using tm_left_until_def transforms_def transits_tm_left_until[OF assms] by simp text \ The most common case is $H = \{\triangleright\}$, which means the Turing machine moves the tape head left to the closest start symbol. On clean tapes it moves the tape head to the leftmost cell of the tape. \ definition tm_start :: "tapeidx \ machine" where "tm_start \ tm_left_until {1}" lemma tm_start_tm: assumes "k \ 2" and "G \ 4" shows "turing_machine k G (tm_start j)" unfolding tm_start_def using assms tm_left_until_tm by simp lemma transforms_tm_start: assumes "j < length tps" and "clean_tape (tps ! j)" shows "transforms (tm_start j) tps (Suc (tps :#: j)) (tps[j := tps ! j |#=| 0])" using tm_start_def assms transforms_tm_left_until begin_tape_def clean_tape_def by (metis insertCI singletonD) lemma transforms_tm_startI [transforms_intros]: assumes "j < length tps" and "clean_tape (tps ! j)" and "t = Suc (tps :#: j)" and "tps' = tps[j := tps ! j |#=| 0]" shows "transforms (tm_start j) tps t tps'" using transforms_tm_start assms by simp text \ The next Turing machine is the first instance in which we use the $;;$ operator with concrete Turing machines. It is also the first time we use the proof method @{method tform} for @{const transforms}. The TM performs a ``carriage return'' on a clean tape, that is, it moves to the first non-start symbol. \ definition tm_cr :: "tapeidx \ machine" where "tm_cr j \ tm_start j ;; tm_right j" lemma tm_cr_tm: "k \ 2 \ G \ 4 \ turing_machine k G (tm_cr j)" using turing_machine_sequential_turing_machine by (simp add: tm_cr_def tm_right_tm tm_start_tm) lemma transforms_tm_crI [transforms_intros]: assumes "j < length tps" and "clean_tape (tps ! j)" and "t = tps :#: j + 2" and "tps' = tps[j := tps ! j |#=| 1]" shows "transforms (tm_cr j) tps t tps'" unfolding tm_cr_def by (tform tps: assms) subsection \Erasing a tape\ text \ The next Turing machine overwrites all but the start symbol with blanks. It first performs a carriage return and then writes blanks until it reaches a blank. This only works as intended if there are no gaps, that is, blanks between non-blank symbols. \ definition tm_erase :: "tapeidx \ machine" where "tm_erase j \ tm_cr j ;; tm_const_until j j {\} \" lemma tm_erase_tm: "G \ 4 \ 0 < j \ j < k \ turing_machine k G (tm_erase j)" unfolding tm_erase_def using tm_cr_tm tm_const_until_tm by simp lemma transforms_tm_eraseI [transforms_intros]: assumes "j < length tps" and "proper_symbols zs" and "tps ::: j = \zs\" and "t = tps :#: j + length zs + 3" and "tps' = tps[j := (\[]\, Suc (length zs))]" shows "transforms (tm_erase j) tps t tps'" unfolding tm_erase_def proof (tform tps: assms time: assms(4)) show "clean_tape (tps ! j)" using assms contents_clean_tape' by simp show "rneigh (tps[j := tps ! j |#=| 1] ! j) {\} (length zs)" using assms contents_clean_tape' by (intro rneighI) auto show "tps' = tps [j := tps ! j |#=| 1, j := tps[j := tps ! j |#=| 1] ! j |+| length zs, j := constplant (tps[j := tps ! j |#=| 1] ! j) \ (length zs)]" proof - have "(\[]\, Suc (length zs)) = constplant (\zs\, Suc 0) \ (length zs)" using transplant_def contents_def by auto then show ?thesis using assms by simp qed qed text \ The next TM returns to the leftmost blank symbol after erasing the tape. \ definition tm_erase_cr :: "tapeidx \ machine" where "tm_erase_cr j \ tm_erase j ;; tm_cr j" lemma tm_erase_cr_tm: assumes "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_erase_cr j)" using tm_erase_cr_def tm_cr_tm tm_erase_tm assms by simp lemma transforms_tm_erase_crI [transforms_intros]: assumes "j < length tps" and "proper_symbols zs" and "tps ::: j = \zs\" and "t = tps :#: j + 2 * length zs + 6" and "tps' = tps[j := (\[]\, 1)]" shows "transforms (tm_erase_cr j) tps t tps'" unfolding tm_erase_cr_def by (tform tps: assms time: assms(4)) subsection \Writing a symbol sequence\ text \ The Turing machine in this section writes a hard-coded symbol sequence to a tape. It is like @{const tm_write_repeat} except with an arbitrary symbol sequence. \ fun tm_print :: "tapeidx \ symbol list \ machine" where "tm_print j [] = []" | "tm_print j (z # zs) = tm_char j z ;; tm_print j zs" lemma tm_print_tm: assumes "0 < j" and "j < k" and "G \ 4" and "\i The result of writing the symbols @{term zs} to a tape @{term tp}: \ definition inscribe :: "tape \ symbol list \ tape" where "inscribe tp zs \ (\i. if snd tp \ i \ i < snd tp + length zs then zs ! (i - snd tp) else fst tp i, snd tp + length zs)" lemma inscribe_Nil: "inscribe tp [] = tp" proof - have "(\i. if snd tp \ i \ i < snd tp then [] ! (i - snd tp) else fst tp i) = fst tp" by auto then show ?thesis unfolding inscribe_def by simp qed lemma inscribe_Cons: "inscribe ((fst tp)(snd tp := z), Suc (snd tp)) zs = inscribe tp (z # zs)" using inscribe_def by auto lemma inscribe_contents: "inscribe (\ys\, Suc (length ys)) zs = (\ys @ zs\, Suc (length ys + length zs))" (is "?lhs = ?rhs") proof show "snd ?lhs = snd ?rhs" using inscribe_def contents_def by simp show "fst ?lhs = fst ?rhs" proof fix i :: nat consider "i = 0" | "0 < i \ i < Suc (length ys)" | "Suc (length ys) \ i \ i < Suc (length ys + length zs)" | "Suc (length ys + length zs) \ i" by linarith then show "fst ?lhs i = fst ?rhs i" proof (cases) case 1 then show ?thesis using inscribe_def contents_def by simp next case 2 then have "fst ?lhs i = \ys\ i" using inscribe_def by simp then have lhs: "fst ?lhs i = ys ! (i - 1)" using 2 contents_def by simp have "fst ?rhs i = (ys @ zs) ! (i - 1)" using 2 contents_def by simp then have "fst ?rhs i = ys ! (i - 1)" using 2 by (metis Suc_diff_1 not_less_eq nth_append) then show ?thesis using lhs by simp next case 3 then show ?thesis using contents_def inscribe_def by (smt (verit, del_insts) One_nat_def add.commute diff_Suc_eq_diff_pred fst_conv length_append less_Suc0 less_Suc_eq_le less_diff_conv2 nat.simps(3) not_less_eq nth_append plus_1_eq_Suc snd_conv) next case 4 then show ?thesis using contents_def inscribe_def by simp qed qed qed lemma inscribe_contents_Nil: "inscribe (\[]\, Suc 0) zs = (\zs\, Suc (length zs))" using inscribe_def contents_def by auto lemma transforms_tm_print: assumes "j < length tps" shows "transforms (tm_print j zs) tps (length zs) (tps[j := inscribe (tps ! j) zs])" using assms proof (induction zs arbitrary: tps) case Nil then show ?case using inscribe_Nil transforms_Nil by simp next case (Cons z zs) have "transforms (tm_char j z ;; tm_print j zs) tps (length (z # zs)) (tps[j := inscribe (tps ! j) (z # zs)])" proof (tform tps: Cons) let ?tps = "tps[j := tps ! j |:=| z |+| 1]" have "transforms (tm_print j zs) ?tps (length zs) (?tps[j := inscribe (?tps ! j) zs])" using Cons by (metis length_list_update) moreover have "(?tps[j := inscribe (?tps ! j) zs]) = (tps[j := inscribe (tps ! j) (z # zs)])" using inscribe_Cons Cons.prems by simp ultimately show "transforms (tm_print j zs) ?tps (length zs) (tps[j := inscribe (tps ! j) (z # zs)])" by simp qed then show "transforms (tm_print j (z # zs)) tps (length (z # zs)) (tps[j := inscribe (tps ! j) (z # zs)])" by simp qed lemma transforms_tm_printI [transforms_intros]: assumes "j < length tps" and "tps' = (tps[j := inscribe (tps ! j) zs])" shows "transforms (tm_print j zs) tps (length zs) tps'" using assms transforms_tm_print by simp subsection \Setting the tape contents to a symbol sequence\ text \ The following Turing machine erases the tape, then prints a hard-coded symbol sequence, and then performs a carriage return. It thus sets the tape contents to the symbol sequence. \ definition tm_set :: "tapeidx \ symbol list \ machine" where "tm_set j zs \ tm_erase_cr j ;; tm_print j zs ;; tm_cr j" lemma tm_set_tm: assumes "0 < j" and "j < k" and "G \ 4" and "\iys\" and "t = 8 + tps :#: j + 2 * length ys + Suc (2 * length zs)" and "tps' = tps[j := (\zs\, 1)]" shows "transforms (tm_set j zs) tps t tps'" unfolding tm_set_def proof (tform tps: assms(1-5)) show "clean_tape (tps[j := (\[]\, 1), j := inscribe (tps[j := (\[]\, 1)] ! j) zs] ! j)" using assms inscribe_contents_Nil clean_contents_proper[OF assms(4)] by simp show "tps' = tps [j := (\[]\, 1), j := inscribe (tps[j := (\[]\, 1)] ! j) zs, j := tps[j := (\[]\, 1), j := inscribe (tps[j := (\[]\, 1)] ! j) zs] ! j |#=| 1]" using assms inscribe_def clean_tape_def assms contents_def inscribe_contents_Nil by simp show "t = tps :#: j + 2 * length ys + 6 + length zs + (tps[j := (\[]\, 1), j := inscribe (tps[j := (\[]\, 1)] ! j) zs] :#: j + 2)" using assms inscribe_def by simp qed subsection \Comparing two tapes\label{s:tm-elementary-comparing}\ text \ The next Turing machine compares the contents of two tapes $j_1$ and $j_2$ and writes to tape $j_3$ either a @{text \} or a @{text \} depending on whether the tapes are equal or not. The next command does all the work. It scans both tapes left to right and halts if it encounters a blank on both tapes, which means the tapes are equal, or two different symbols, which means the tapes are unequal. It only works for contents without blanks. \ definition cmd_cmp :: "tapeidx \ tapeidx \ tapeidx \ command" where "cmd_cmp j1 j2 j3 rs \ if rs ! j1 \ rs ! j2 then (1, map (\i. (if i = j3 then \ else rs ! i, Stay)) [0.. \ rs ! j2 = \ then (1, map (\i. (if i = j3 then \ else rs ! i, Stay)) [0..i. (rs ! i, if i = j1 \ i = j2 then Right else Stay)) [0.. tps :.: j2" shows "sem (cmd_cmp j1 j2 j3) (0, tps) = (1, tps[j3 := tps ! j3 |:=| \])" unfolding cmd_cmp_def using assms tapes_at_read' act_Stay read_length by (intro semI) auto lemma sem_cmd_cmp2: assumes "length tps = k" and "j1 < k" and "j2 < k" and "j3 < k" and "tps :.: j1 = tps :.: j2" and "tps :.: j1 = \ \ tps :.: j2 = \" shows "sem (cmd_cmp j1 j2 j3) (0, tps) = (1, tps[j3 := tps ! j3 |:=| \])" unfolding cmd_cmp_def using assms tapes_at_read' act_Stay read_length by (intro semI) auto lemma sem_cmd_cmp3: assumes "length tps = k" and "j2 \ j3" and "j1 \ j3" and "j1 < k" and "j2 < k" and "j3 < k" and "tps :.: j1 = tps :.: j2" and "tps :.: j1 \ \ \ tps :.: j2 \ \" shows "sem (cmd_cmp j1 j2 j3) (0, tps) = (0, tps[j1 := tps ! j1 |+| 1, j2 := tps ! j2 |+| 1])" proof (rule semI) show "proper_command k (cmd_cmp j1 j2 j3)" using cmd_cmp_def by simp show "length tps = k" using assms(1) . show "length (tps[j1 := tps ! j1 |+| 1, j2 := tps ! j2 |+| 1]) = k" using assms(1) by simp show "fst ((cmd_cmp j1 j2 j3) (read tps)) = 0" unfolding cmd_cmp_def using assms tapes_at_read' by simp show "act (cmd_cmp j1 j2 j3 (read tps) [!] j) (tps ! j) = tps[j1 := tps ! j1 |+| 1, j2 := tps ! j2 |+| 1] ! j" if "j < k" for j unfolding cmd_cmp_def using assms tapes_at_read' that act_Stay act_Right read_length by (cases "j1 = j2") simp_all qed definition tm_cmp :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_cmp j1 j2 j3 \ [cmd_cmp j1 j2 j3]" lemma tm_cmp_tm: assumes "k \ 2" and "j3 > 0" and "G \ 4" shows "turing_machine k G (tm_cmp j1 j2 j3)" unfolding tm_cmp_def cmd_cmp_def using assms turing_machine_def by auto lemma exe_cmd_cmp1: assumes "length tps = k" and "j1 < k" and "j2 < k" and "j3 < k" and "tps :.: j1 \ tps :.: j2" shows "exe (tm_cmp j1 j2 j3) (0, tps) = (1, tps[j3 := tps ! j3 |:=| \])" using tm_cmp_def assms exe_lt_length sem_cmd_cmp1 by simp lemma exe_cmd_cmp2: assumes "length tps = k" and "j1 < k" and "j2 < k" and "j3 < k" and "tps :.: j1 = tps :.: j2" and "tps :.: j1 = \ \ tps :.: j2 = \" shows "exe (tm_cmp j1 j2 j3) (0, tps) = (1, tps[j3 := tps ! j3 |:=| \])" using tm_cmp_def assms exe_lt_length sem_cmd_cmp2 by simp lemma exe_cmd_cmp3: assumes "length tps = k" and "j2 \ j3" and "j1 \ j3" and "j1 < k" and "j2 < k" and "j3 < k" and "tps :.: j1 = tps :.: j2" and "tps :.: j1 \ \ \ tps :.: j2 \ \" shows "exe (tm_cmp j1 j2 j3) (0, tps) = (0, tps[j1 := tps ! j1 |+| 1, j2 := tps ! j2 |+| 1])" using tm_cmp_def assms exe_lt_length sem_cmd_cmp3 by simp lemma execute_tm_cmp_eq: fixes tps :: "tape list" assumes "length tps = k" and "j2 \ j3" and "j1 \ j3" and "j1 < k" and "j2 < k" and "j3 < k" and "proper_symbols xs" and "tps ! j1 = (\xs\, 1)" and "tps ! j2 = (\xs\, 1)" shows "execute (tm_cmp j1 j2 j3) (0, tps) (Suc (length xs)) = (1, tps[j1 := tps ! j1 |+| length xs, j2 := tps ! j2 |+| length xs, j3 := tps ! j3 |:=| \])" proof - have "execute (tm_cmp j1 j2 j3) (0, tps) t = (0, tps[j1 := tps ! j1 |+| t, j2 := tps ! j2 |+| t])" if "t \ length xs" for t using that proof (induction t) case 0 then show ?case by simp next case (Suc t) then have t_less: "t < length xs" by simp have "execute (tm_cmp j1 j2 j3) (0, tps) (Suc t) = exe (tm_cmp j1 j2 j3) (execute (tm_cmp j1 j2 j3) (0, tps) t)" by simp also have "... = exe (tm_cmp j1 j2 j3) (0, tps[j1 := tps ! j1 |+| t, j2 := tps ! j2 |+| t])" (is "_ = exe _ (0, ?tps)") using Suc by simp also have "... = (0, ?tps[j1 := ?tps ! j1 |+| 1, j2 := ?tps ! j2 |+| 1])" proof - have 1: "?tps :.: j1 = xs ! t" using assms(1,2,4,8) t_less Suc.prems contents_inbounds by (metis (no_types, lifting) diff_Suc_1 fst_conv length_list_update nth_list_update_eq nth_list_update_neq plus_1_eq_Suc snd_conv zero_less_Suc) moreover have 2: "?tps :.: j2 = xs ! t" using t_less assms(1,5,9) by simp ultimately have "?tps :.: j1 = ?tps :.: j2" by simp moreover have "?tps :.: j1 \ 0 \ ?tps :.: j2 \ 0" using 1 2 assms(7) t_less by auto moreover have "length ?tps = k" using assms(1) by simp ultimately show ?thesis using assms exe_cmd_cmp3 by blast qed also have "... = (0, tps[j1 := tps ! j1 |+| Suc t, j2 := tps ! j2 |+| Suc t])" using assms - by (smt (z3) Suc_eq_plus1 add.commute fst_conv list_update_overwrite list_update_swap + by (smt (verit) Suc_eq_plus1 add.commute fst_conv list_update_overwrite list_update_swap nth_list_update_eq nth_list_update_neq snd_conv) finally show ?case by simp qed then have "execute (tm_cmp j1 j2 j3) (0, tps) (length xs) = (0, tps[j1 := tps ! j1 |+| length xs, j2 := tps ! j2 |+| length xs])" by simp then have "execute (tm_cmp j1 j2 j3) (0, tps) (Suc (length xs)) = exe (tm_cmp j1 j2 j3) (0, tps[j1 := tps ! j1 |+| length xs, j2 := tps ! j2 |+| length xs])" (is "_ = exe _ (0, ?tps)") by simp also have "... = (1, ?tps[j3 := ?tps ! j3 |:=| \])" proof - have 1: "?tps :.: j1 = 0" using assms(1,4,8) contents_outofbounds by (metis fst_conv length_list_update lessI nth_list_update_eq nth_list_update_neq plus_1_eq_Suc snd_conv) moreover have 2: "?tps :.: j2 = 0" using assms(1,5,9) by simp ultimately have "?tps :.: j1 = ?tps :.: j2" "?tps :.: j1 = \ \ ?tps :.: j2 = \" by simp_all moreover have "length ?tps = k" using assms(1) by simp ultimately show ?thesis using assms exe_cmd_cmp2 by blast qed also have "... = (1, tps[j1 := tps ! j1 |+| length xs, j2 := tps ! j2 |+| length xs, j3 := tps ! j3 |:=| \])" using assms by simp finally show ?thesis . qed lemma ex_contents_neq: assumes "proper_symbols xs" and "proper_symbols ys" and "xs \ ys" shows "\m. m \ Suc (min (length xs) (length ys)) \ \xs\ m \ \ys\ m" proof - consider "length xs < length ys" | "length xs = length ys" | "length xs > length ys" by linarith then show ?thesis proof (cases) case 1 let ?m = "length xs" have "\xs\ (Suc ?m) = \" by simp moreover have "\ys\ (Suc ?m) \ \" using 1 assms(2) by (simp add: proper_symbols_ne0) ultimately show ?thesis using 1 by auto next case 2 then have "\i ys ! i" using assms by (meson list_eq_iff_nth_eq) then show ?thesis using contents_def 2 by auto next case 3 let ?m = "length ys" have "\ys\ (Suc ?m) = \" by simp moreover have "\xs\ (Suc ?m) \ \" using 3 assms(1) by (simp add: proper_symbols_ne0) ultimately show ?thesis using 3 by auto qed qed lemma execute_tm_cmp_neq: fixes tps :: "tape list" assumes "length tps = k" and "j1 \ j2" and "j2 \ j3" and "j1 \ j3" and "j1 < k" and "j2 < k" and "j3 < k" and "proper_symbols xs" and "proper_symbols ys" and "xs \ ys" and "tps ! j1 = (\xs\, 1)" and "tps ! j2 = (\ys\, 1)" and "m = (LEAST m. m \ Suc (min (length xs) (length ys)) \ \xs\ m \ \ys\ m)" shows "execute (tm_cmp j1 j2 j3) (0, tps) m = (1, tps[j1 := tps ! j1 |+| (m - 1), j2 := tps ! j2 |+| (m - 1), j3 := tps ! j3 |:=| \])" proof - have neq: "\xs\ m \ \ys\ m" using ex_contents_neq[OF assms(8-10)] assms(13) by (metis (mono_tags, lifting) LeastI_ex) have eq: "\xs\ i = \ys\ i" if "i < m" for i - using ex_contents_neq[OF assms(8-10)] assms(13) not_less_Least that by (smt (z3) Least_le le_trans less_imp_le_nat) + using ex_contents_neq[OF assms(8-10)] assms(13) not_less_Least that by (smt (verit) Least_le le_trans less_imp_le_nat) have "m > 0" using neq contents_def gr0I by metis have "execute (tm_cmp j1 j2 j3) (0, tps) t = (0, tps[j1 := tps ! j1 |+| t, j2 := tps ! j2 |+| t])" if "t < m" for t using that proof (induction t) case 0 then show ?case by simp next case (Suc t) have "execute (tm_cmp j1 j2 j3) (0, tps) (Suc t) = exe (tm_cmp j1 j2 j3) (execute (tm_cmp j1 j2 j3) (0, tps) t)" by simp also have "... = exe (tm_cmp j1 j2 j3) (0, tps[j1 := tps ! j1 |+| t, j2 := tps ! j2 |+| t])" (is "_ = exe _ (0, ?tps)") using Suc by simp also have "... = (0, ?tps[j1 := ?tps ! j1 |+| 1, j2 := ?tps ! j2 |+| 1])" proof - have 1: "?tps :.: j1 = \xs\ (Suc t)" using assms(1,2,5,11) by simp moreover have 2: "?tps :.: j2 = \ys\ (Suc t)" using assms(1,6,12) by simp ultimately have "?tps :.: j1 = ?tps :.: j2" using Suc eq by simp moreover from this have "?tps :.: j1 \ \ \ ?tps :.: j2 \ \" using 1 2 assms neq Suc.prems contents_def - by (smt (z3) Suc_leI Suc_le_lessD Suc_lessD diff_Suc_1 le_trans less_nat_zero_code zero_less_Suc) + by (smt (verit) Suc_leI Suc_le_lessD Suc_lessD diff_Suc_1 le_trans less_nat_zero_code zero_less_Suc) moreover have "length ?tps = k" using assms(1) by simp ultimately show ?thesis using assms exe_cmd_cmp3 by blast qed also have "... = (0, tps[j1 := tps ! j1 |+| Suc t, j2 := tps ! j2 |+| Suc t])" using assms - by (smt (z3) Suc_eq_plus1 add.commute fst_conv list_update_overwrite list_update_swap + by (smt (verit) Suc_eq_plus1 add.commute fst_conv list_update_overwrite list_update_swap nth_list_update_eq nth_list_update_neq snd_conv) finally show ?case by simp qed then have "execute (tm_cmp j1 j2 j3) (0, tps) (m - 1) = (0, tps[j1 := tps ! j1 |+| (m - 1), j2 := tps ! j2 |+| (m - 1)])" using `m > 0` by simp then have "execute (tm_cmp j1 j2 j3) (0, tps) m = exe (tm_cmp j1 j2 j3) (0, tps[j1 := tps ! j1 |+| (m - 1), j2 := tps ! j2 |+| (m - 1)])" using `m > 0` by (metis contents_at_0 diff_Suc_1 execute.elims neq) then show "execute (tm_cmp j1 j2 j3) (0, tps) m = (1, tps[j1 := tps ! j1 |+| (m - 1), j2 := tps ! j2 |+| (m - 1), j3 := tps ! j3 |:=| \])" using exe_cmd_cmp1 assms \0 < m\ - by (smt (z3) One_nat_def Suc_diff_Suc diff_zero fst_conv length_list_update neq nth_list_update_eq + by (smt (verit) One_nat_def Suc_diff_Suc diff_zero fst_conv length_list_update neq nth_list_update_eq nth_list_update_neq plus_1_eq_Suc snd_conv) qed lemma transforms_tm_cmpI [transforms_intros]: fixes tps :: "tape list" assumes "length tps = k" and "j1 \ j2" and "j2 \ j3" and "j1 \ j3" and "j1 < k" and "j2 < k" and "j3 < k" and "proper_symbols xs" and "proper_symbols ys" and "tps ! j1 = (\xs\, 1)" and "tps ! j2 = (\ys\, 1)" and "t = Suc (min (length xs) (length ys))" and "b = (if xs = ys then \ else \)" and "m = (if xs = ys then Suc (length xs) else (LEAST m. m \ Suc (min (length xs) (length ys)) \ \xs\ m \ \ys\ m))" and "tps' = tps[j1 := (\xs\, m), j2 := (\ys\, m), j3 := tps ! j3 |:=| b]" shows "transforms (tm_cmp j1 j2 j3) tps t tps'" proof (cases "xs = ys") case True then have m: "m = Suc (length xs)" using assms(14) by simp have "execute (tm_cmp j1 j2 j3) (0, tps) (Suc (length xs)) = (1, tps[j1 := tps ! j1 |+| length xs, j2 := tps ! j2 |+| length xs, j3 := tps ! j3 |:=| \])" using execute_tm_cmp_eq assms True by blast then have "execute (tm_cmp j1 j2 j3) (0, tps) m = (1, tps[j1 := tps ! j1 |+| (m - 1), j2 := tps ! j2 |+| (m - 1), j3 := tps ! j3 |:=| b])" using m assms(13) True diff_Suc_1 by simp moreover have "m \ t" using True assms(12) m by simp ultimately show ?thesis using transitsI tm_cmp_def transforms_def assms True by (metis (no_types, lifting) One_nat_def add.commute diff_Suc_1 fst_conv list.size(3) list.size(4) plus_1_eq_Suc snd_conv) next case False then have m: "m = (LEAST m. m \ Suc (min (length xs) (length ys)) \ \xs\ m \ \ys\ m)" using assms(14) by simp then have "execute (tm_cmp j1 j2 j3) (0, tps) m = (1, tps[j1 := tps ! j1 |+| (m - 1), j2 := tps ! j2 |+| (m - 1), j3 := tps ! j3 |:=| \])" using False assms execute_tm_cmp_neq by blast then have "execute (tm_cmp j1 j2 j3) (0, tps) m = (1, tps[j1 := tps ! j1 |+| (m - 1), j2 := tps ! j2 |+| (m - 1), j3 := tps ! j3 |:=| b])" using False by (simp add: assms(13)) moreover have "m \ t" using ex_contents_neq[OF assms(8,9)] False assms(12) m by (metis (mono_tags, lifting) LeastI) ultimately show ?thesis using transitsI tm_cmp_def transforms_def assms False by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 Suc_pred add.commute execute.simps(1) fst_eqD list.size(3) list.size(4) not_gr0 numeral_One snd_conv zero_neq_numeral) qed text \ The next Turing machine extends @{const tm_cmp} by a carriage return on tapes $j_1$ and $j_2$, ensuring that the next command finds the tape heads in a well-specified position. This makes the TM easier to reuse. \ definition tm_equals :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_equals j1 j2 j3 \ tm_cmp j1 j2 j3 ;; tm_cr j1 ;; tm_cr j2" lemma tm_equals_tm: assumes "k \ 2" and "j3 > 0" and "G \ 4" shows "turing_machine k G (tm_equals j1 j2 j3)" unfolding tm_equals_def using tm_cmp_tm tm_cr_tm assms by simp text \ We analyze the behavior of @{const tm_equals} inside a locale. This is how we will typically proceed for Turing machines that are composed of more than two TMs. The locale is parameterized by the TM's parameters, which in the present case means the three tape indices $j_1$, $j_2$, and $j_3$. Inside the locale the TM is decomposed such that proofs of @{const transforms} only involve two TMs combined by one of the three control structures (sequence, branch, loop). In the current example we have three TMs named @{term tm1}, @{term tm2}, @{term tm3}, where @{term tm3} is just @{const tm_equals}. Furthermore there will be lemmas \emph{tm1}, \emph{tm2}, \emph{tm3} describing, in terms of @{const transforms}, the behavior of the respective TMs. For this we define three tape lists @{term tps1}, @{term tps2}, @{term tps3}. This naming scheme creates many name clashes for things that only have a single use. That is the reason for the encapsulation in a locale. Afterwards this locale is interpreted, just once in lemma~@{text transforms_tm_equalsI}, to prove the semantics and running time of @{const tm_equals}. \null \ locale turing_machine_equals = fixes j1 j2 j3 :: tapeidx begin definition "tm1 \ tm_cmp j1 j2 j3" definition "tm2 \ tm1 ;; tm_cr j1" definition "tm3 \ tm2 ;; tm_cr j2" lemma tm3_eq_tm_equals: "tm3 = tm_equals j1 j2 j3" unfolding tm3_def tm2_def tm1_def tm_equals_def by simp context fixes tps0 :: "tape list" and k t b :: nat and xs ys :: "symbol list" assumes jk [simp]: "length tps0 = k" "j1 \ j2" "j2 \ j3" "j1 \ j3" "j1 < k" "j2 < k" "j3 < k" and proper: "proper_symbols xs" "proper_symbols ys" and t: "t = Suc (min (length xs) (length ys))" and b: "b = (if xs = ys then 3 else 0)" assumes tps0: "tps0 ! j1 = (\xs\, 1)" "tps0 ! j2 = (\ys\, 1)" begin definition "m \ (if xs = ys then Suc (length xs) else (LEAST m. m \ Suc (min (length xs) (length ys)) \ \xs\ m \ \ys\ m))" lemma m_gr_0: "m > 0" proof - have "\xs\ m \ \ys\ m" if "xs \ ys" using ex_contents_neq LeastI_ex m_def proper that by (metis (mono_tags, lifting)) then show "m > 0" using m_def by (metis contents_at_0 gr0I less_Suc_eq_0_disj) qed lemma m_le_t: "m \ t" proof (cases "xs = ys") case True then show ?thesis using t m_def by simp next case False then have "m \ Suc (min (length xs) (length ys))" using ex_contents_neq False proper m_def by (metis (mono_tags, lifting) LeastI_ex) then show ?thesis using t by simp qed definition "tps1 \ tps0[j1 := (\xs\, m), j2 := (\ys\, m), j3 := tps0 ! j3 |:=| b]" lemma tm1 [transforms_intros]: "transforms tm1 tps0 t tps1" unfolding tm1_def proof (tform tps: tps0 tps1_def m_def b time: t) show "proper_symbols xs" "proper_symbols ys" using proper by simp_all qed definition "tps2 \ tps0[j1 := (\xs\, 1), j2 := (\ys\, m), j3 := tps0 ! j3 |:=| b]" lemma tm2: assumes "ttt = t + m + 2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps1_def) show "clean_tape (tps1 ! j1)" using tps1_def clean_contents_proper jk proper(1) by (metis nth_list_update_eq nth_list_update_neq) show "ttt = t + (tps1 :#: j1 + 2)" using tps1_def tps0 jk assms by (metis (no_types, lifting) ab_semigroup_add_class.add_ac(1) nth_list_update_eq nth_list_update_neq snd_conv) show "tps2 = tps1[j1 := tps1 ! j1 |#=| 1]" unfolding tps2_def tps1_def by (simp add: list_update_swap[of j1]) qed lemma tm2' [transforms_intros]: "transforms tm2 tps0 (2 * t + 2) tps2" using m_le_t tm2 transforms_monotone by simp definition "tps3 \ tps0[j1 := (\xs\, 1), j2 := (\ys\, 1), j3 := tps0 ! j3 |:=| b]" lemma tm3: assumes "ttt = 2 * t + m + 4" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps2_def tps3_def) have *: "tps2 ! j2 = (\ys\, m)" using tps2_def by (simp add: nth_list_update_neq') then show "clean_tape (tps2 ! j2)" using clean_contents_proper proper(2) by simp show "ttt = 2 * t + 2 + (tps2 :#: j2 + 2)" using assms * by simp show "tps3 = tps2[j2 := tps2 ! j2 |#=| 1]" unfolding tps3_def tps2_def by (simp add: list_update_swap[of j2]) qed definition "tps3' \ tps0[j3 := tps0 ! j3 |:=| b]" lemma tm3': "transforms tm3 tps0 (3 * min (length xs) (length ys) + 7) tps3'" proof - have "tps3' = tps3" using tps3'_def tps3_def tps0 jk by (metis list_update_id) then show ?thesis using m_le_t tm3 transforms_monotone t by simp qed end (* context tps0 k xs ys *) end (* locale turing_machine_equals *) lemma transforms_tm_equalsI [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and k :: nat and xs ys :: "symbol list" and b :: symbol assumes "length tps = k" "j1 \ j2" "j2 \ j3" "j1 \ j3" "j1 < k" "j2 < k" "j3 < k" and "proper_symbols xs" "proper_symbols ys" and "b = (if xs = ys then \ else \)" assumes "tps ! j1 = (\xs\, 1)" "tps ! j2 = (\ys\, 1)" assumes "ttt = 3 * min (length xs) (length ys) + 7" assumes "tps' = tps [j3 := tps ! j3 |:=| b]" shows "transforms (tm_equals j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_equals j1 j2 j3 . show ?thesis using assms loc.tm3' loc.tm3_eq_tm_equals loc.tps3'_def by simp qed subsection \Computing the identity function\ text \ In order to compute the identity function, a Turing machine can just copy the input tape to the output tape: \ definition tm_id :: machine where "tm_id \ tm_cp_until 0 1 {\}" lemma tm_id_tm: assumes "1 < k" and "G \ 4" shows "turing_machine k G tm_id" unfolding tm_id_def using assms tm_cp_until_tm by simp lemma transforms_tm_idI: fixes zs :: "symbol list" and k :: nat and tps :: "tape list" assumes "1 < k" and "proper_symbols zs" and "tps = snd (start_config k zs)" and "tps' = tps[0 := (\zs\, (Suc (length zs))), 1 := (\zs\, (Suc (length zs)))]" shows "transforms tm_id tps (Suc (Suc (length zs))) tps'" proof - let ?n = "Suc (length zs)" define tps2 where "tps2 = tps[0 := tps ! 0 |+| (Suc (length zs)), 1 := implant (tps ! 0) (tps ! 1) (Suc (length zs))]" have 1: "rneigh (tps ! 0) {\} ?n" proof (rule rneighI) show "(tps ::: 0) (tps :#: 0 + Suc (length zs)) \ {\}" using start_config2 start_config3 assms by (simp add: start_config_def) show "\n'. n' < Suc (length zs) \ (tps ::: 0) (tps :#: 0 + n') \ {\}" using start_config2 start_config3 start_config_pos assms by (metis One_nat_def Suc_lessD add_cancel_right_left diff_Suc_1 less_Suc_eq_0_disj less_Suc_eq_le not_one_less_zero singletonD) qed have 2: "length tps = k" using assms(1,3) by (simp add: start_config_length) have **: "transforms tm_id tps (Suc ?n) tps2" unfolding tm_id_def using transforms_tm_cp_untilI[OF _ assms(1) 2 1 _ tps2_def] assms(1) by simp have 0: "tps ! 0 = (\zs\, 0)" using assms start_config_def contents_def by auto moreover have "tps ! 1 = (\[]\, 0)" using assms start_config_def contents_def by auto moreover have "implant (\zs\, 0) (\[]\, 0) ?n = (\zs\, ?n)" by (rule implantI''') simp_all ultimately have "implant (tps ! 0) (tps ! 1) ?n = (\zs\, ?n)" by simp then have "tps2 = tps[0 := tps ! 0 |+| ?n, 1 := (\zs\, ?n)]" using tps2_def by simp then have "tps2 = tps[0 := (\zs\, ?n), 1 := (\zs\, ?n)]" using 0 by simp then have "tps2 = tps'" using assms(4) by simp then show ?thesis using ** by simp qed text \ The identity function is computable with a time bound of $n + 2$. \ lemma computes_id: "computes_in_time 2 tm_id id (\n. Suc (Suc n))" proof fix x :: string let ?zs = "string_to_symbols x" let ?start = "snd (start_config 2 ?zs)" let ?T = "\n. Suc (Suc n)" let ?tps = "?start[0 := (\?zs\, (Suc (length ?zs))), 1 := (\?zs\, (Suc (length ?zs)))]" have "proper_symbols ?zs" by simp then have "transforms tm_id ?start (Suc (Suc (length ?zs))) ?tps" using transforms_tm_idI One_nat_def less_2_cases_iff by blast then have "transforms tm_id ?start (?T (length x)) ?tps" by simp moreover have "?tps ::: 1 = string_to_contents (id x)" by (auto simp add: start_config_length) ultimately show "\tps. tps ::: 1 = string_to_contents (id x) \ transforms tm_id ?start (?T (length x)) tps" by auto qed end diff --git a/thys/Cook_Levin/Lists_Lists.thy b/thys/Cook_Levin/Lists_Lists.thy --- a/thys/Cook_Levin/Lists_Lists.thy +++ b/thys/Cook_Levin/Lists_Lists.thy @@ -1,3898 +1,3898 @@ section \Lists of numbers\label{s:tm-numlist}\ theory Lists_Lists imports Arithmetic begin text \ In the previous section we defined a representation for numbers over the binary alphabet @{text "{\,\}"}. In this section we first introduce a representation of lists of numbers as symbol sequences over the alphabet @{text "{\,\,\}"}. Then we define Turing machines for some operations over such lists. As with the arithmetic operations, Turing machines implementing the operations on lists usually expect the tape heads of the operands to be in position~1 and guarantee to place the tape heads of the result in position~1. The exception are Turing machines for iterating over lists; such TMs move the tape head to the next list element. A tape containing such representations corresponds to a variable of type @{typ "nat list"}. A tape in the start configuration corresponds to the empty list of numbers. \ subsection \Representation as symbol sequence\label{s:tm-numlist-repr}\ text \ The obvious idea for representing a list of numbers is to write them one after another separated by a fresh symbol, such as @{text "\"}. However since we represent the number~0 by the empty symbol sequence, this would result in both the empty list and the list containing only the number~0 to be represented by the same symbol sequence, namely the empty one. To prevent this we use the symbol @{text "\"} not as a separator but as a terminator; that is, we append it to every number. Consequently the empty symbol sequence represents the empty list, whereas the list $[0]$ is represented by the symbol sequence @{text "\"}. As another example, the list $[14,0,0,7]$ is represented by @{text "\\\\\\\\\\\"}. As a side effect of using terminators instead of separators, the representation of the concatenation of lists is just the concatenation of the representations of the individual lists. Moreover the length of the representation is simply the sum of the individual representation lengths. The number of @{text "\"} symbols equals the number of elements in the list. \ text \ This is how lists of numbers are represented as symbol sequences: \ definition numlist :: "nat list \ symbol list" where "numlist ns \ concat (map (\n. canrepr n @ [\]) ns)" lemma numlist_Nil: "numlist [] = []" using numlist_def by simp proposition "numlist [0] = [\]" using numlist_def by (simp add: canrepr_0) lemma numlist_234: "set (numlist ns) \ {\, \, \}" proof (induction ns) case Nil then show ?case using numlist_def by simp next case (Cons n ns) have "numlist (n # ns) = canrepr n @ [\] @ concat (map (\n. canrepr n @ [\]) ns)" using numlist_def by simp then have "numlist (n # ns) = canrepr n @ [\] @ numlist ns" using numlist_def by simp moreover have "set ([\] @ numlist ns) \ {\, \, \}" using Cons by simp moreover have "set (canrepr n) \ {\, \, \}" using bit_symbols_canrepr by (metis in_set_conv_nth insertCI subsetI) ultimately show ?case by simp qed lemma symbols_lt_numlist: "symbols_lt 5 (numlist ns)" using numlist_234 by (metis empty_iff insert_iff nth_mem numeral_less_iff semiring_norm(68) semiring_norm(76) semiring_norm(79) semiring_norm(80) subset_code(1) verit_comp_simplify1(2)) lemma bit_symbols_prefix_eq: assumes "(x @ [\]) @ xs = (y @ [\]) @ ys" and "bit_symbols x" and "bit_symbols y" shows "x = y" proof - have *: "x @ [\] @ xs = y @ [\] @ ys" (is "?lhs = ?rhs") using assms(1) by simp show "x = y" proof (cases "length x \ length y") case True then have "?lhs ! i = ?rhs ! i" if "i < length x" for i using that * by simp then have eq: "x ! i = y ! i" if "i < length x" for i using that True by (metis Suc_le_eq le_trans nth_append) have "?lhs ! (length x) = \" by (metis Cons_eq_appendI nth_append_length) then have "?rhs ! (length x) = \" using * by metis then have "length y \ length x" using assms(3) by (metis linorder_le_less_linear nth_append numeral_eq_iff semiring_norm(85) semiring_norm(87) semiring_norm(89)) then have "length y = length x" using True by simp then show ?thesis using eq by (simp add: list_eq_iff_nth_eq) next case False then have "?lhs ! i = ?rhs ! i" if "i < length y" for i using that * by simp have "?rhs ! (length y) = \" by (metis Cons_eq_appendI nth_append_length) then have "?lhs ! (length y) = \" using * by metis then have "x ! (length y) = \" using False by (simp add: nth_append) then have False using assms(2) False by (metis linorder_le_less_linear numeral_eq_iff semiring_norm(85) semiring_norm(87) semiring_norm(89)) then show ?thesis by simp qed qed lemma numlist_inj: "numlist ns1 = numlist ns2 \ ns1 = ns2" proof (induction ns1 arbitrary: ns2) case Nil then show ?case using numlist_def by simp next case (Cons n ns1) have 1: "numlist (n # ns1) = (canrepr n @ [\]) @ numlist ns1" using numlist_def by simp then have "numlist ns2 = (canrepr n @ [\]) @ numlist ns1" using Cons by simp then have "ns2 \ []" using numlist_Nil by auto then have 2: "ns2 = hd ns2 # tl ns2" using `ns2 \ []` by simp then have 3: "numlist ns2 = (canrepr (hd ns2) @ [\]) @ numlist (tl ns2)" using numlist_def by (metis concat.simps(2) list.simps(9)) have 4: "hd ns2 = n" using bit_symbols_prefix_eq 1 3 by (metis Cons.prems canrepr bit_symbols_canrepr) then have "numlist ns2 = (canrepr n @ [\]) @ numlist (tl ns2)" using 3 by simp then have "numlist ns1 = numlist (tl ns2)" using 1 by (simp add: Cons.prems) then have "ns1 = tl ns2" using Cons by simp then show ?case using 2 4 by simp qed corollary proper_symbols_numlist: "proper_symbols (numlist ns)" using numlist_234 nth_mem by fastforce text \ The next property would not hold if we used separators between elements instead of terminators after elements. \ lemma numlist_append: "numlist (xs @ ys) = numlist xs @ numlist ys" using numlist_def by simp text \ Like @{const nlength} for numbers, we have @{term nllength} for the length of the representation of a list of numbers. \ definition nllength :: "nat list \ nat" where "nllength ns \ length (numlist ns)" lemma nllength: "nllength ns = (\n\ns. Suc (nlength n))" using nllength_def numlist_def by (induction ns) simp_all lemma nllength_Nil [simp]: "nllength [] = 0" using nllength_def numlist_def by simp lemma length_le_nllength: "length ns \ nllength ns" using nllength sum_list_mono[of ns "\_. 1" "\n. Suc (nlength n)"] sum_list_const[of 1 ns] by simp lemma nllength_le_len_mult_max: fixes N :: nat and ns :: "nat list" assumes "\n\set ns. n \ N" shows "nllength ns \ Suc (nlength N) * length ns" proof - have "nllength ns = (\n\ns. Suc (nlength n))" using nllength by simp moreover have "Suc (nlength n) \ Suc (nlength N)" if "n \ set ns" for n using nlength_mono that assms by simp ultimately have "nllength ns \ (\n\ns. Suc (nlength N))" by (simp add: sum_list_mono) then show "nllength ns \ Suc (nlength N) * length ns" using sum_list_const by metis qed lemma nllength_upt_le_len_mult_max: fixes a b :: nat shows "nllength [a.. Suc (nlength b) * (b - a)" using nllength_le_len_mult_max[of "[a.. Suc (nlength (Max (set ns))) * length ns" using nllength_le_len_mult_max by simp lemma member_le_nllength: "n \ set ns \ nlength n \ nllength ns" using nllength by (induction ns) auto lemma member_le_nllength_1: "n \ set ns \ nlength n \ nllength ns - 1" using nllength by (induction ns) auto lemma nllength_gr_0: "ns \ [] \ 0 < nllength ns" using nllength_def numlist_def by simp lemma nlength_min_le_nllength: "n \ set ns \ m \ set ns \ nlength (min n m) \ nllength ns" using member_le_nllength by (simp add: min_def) lemma take_drop_numlist: assumes "i < length ns" shows "take (Suc (nlength (ns ! i))) (drop (nllength (take i ns)) (numlist ns)) = canrepr (ns ! i) @ [\]" proof - have "numlist ns = numlist (take i ns) @ numlist (drop i ns)" using numlist_append by (metis append_take_drop_id) moreover have "numlist (drop i ns) = numlist [ns ! i] @ numlist (drop (Suc i) ns)" using assms numlist_append by (metis Cons_nth_drop_Suc append_Cons self_append_conv2) ultimately have "numlist ns = numlist (take i ns) @ numlist [ns ! i] @ numlist (drop (Suc i) ns)" by simp then have "drop (nllength (take i ns)) (numlist ns) = numlist [ns ! i] @ numlist (drop (Suc i) ns)" by (simp add: nllength_def) then have "drop (nllength (take i ns)) (numlist ns) = canrepr (ns ! i) @ [\] @ numlist (drop (Suc i) ns)" using numlist_def by simp then show ?thesis by simp qed corollary take_drop_numlist': assumes "i < length ns" shows "take (nlength (ns ! i)) (drop (nllength (take i ns)) (numlist ns)) = canrepr (ns ! i)" using take_drop_numlist[OF assms] by (metis append_assoc append_eq_conv_conj append_take_drop_id) corollary numlist_take_at_term: assumes "i < length ns" shows "numlist ns ! (nllength (take i ns) + nlength (ns ! i)) = \" using assms take_drop_numlist nllength_def numlist_append - by (smt (z3) append_eq_conv_conj append_take_drop_id lessI nth_append_length nth_append_length_plus nth_take) + by (smt (verit) append_eq_conv_conj append_take_drop_id lessI nth_append_length nth_append_length_plus nth_take) lemma numlist_take_at: assumes "i < length ns" and "j < nlength (ns ! i)" shows "numlist ns ! (nllength (take i ns) + j) = canrepr (ns ! i) ! j" proof - have "ns = take i ns @ [ns ! i] @ drop (Suc i) ns" using assms by (metis Cons_eq_appendI append_self_conv2 id_take_nth_drop) then have "numlist ns = (numlist (take i ns) @ numlist [ns ! i]) @ numlist (drop (Suc i) ns)" using numlist_append by (metis append_assoc) moreover have "nllength (take i ns) + j < nllength (take i ns) + nllength [ns ! i]" using assms(2) nllength by simp ultimately have "numlist ns ! (nllength (take i ns) + j) = (numlist (take i ns) @ numlist [ns ! i]) ! (nllength (take i ns) + j)" by (metis length_append nllength_def nth_append) also have "... = numlist [ns ! i] ! j" by (simp add: nllength_def) also have "... = (canrepr (ns ! i) @ [\]) ! j" using numlist_def by simp also have "... = canrepr (ns ! i) ! j" using assms(2) by (simp add: nth_append) finally show ?thesis . qed lemma nllength_take_Suc: assumes "i < length ns" shows "nllength (take i ns) + Suc (nlength (ns ! i)) = nllength (take (Suc i) ns)" proof - have "ns = take i ns @ [ns ! i] @ drop (Suc i) ns" using assms by (metis Cons_eq_appendI append_self_conv2 id_take_nth_drop) then have "numlist ns = numlist (take i ns) @ numlist [ns ! i] @ numlist (drop (Suc i) ns)" using numlist_append by metis then have "nllength ns = nllength (take i ns) + nllength [ns ! i] + nllength (drop (Suc i) ns)" by (simp add: nllength_def) moreover have "nllength ns = nllength (take (Suc i) ns) + nllength (drop (Suc i) ns)" using numlist_append by (metis append_take_drop_id length_append nllength_def) ultimately have "nllength (take (Suc i) ns) = nllength (take i ns) + nllength [ns ! i]" by simp then show ?thesis using nllength by simp qed lemma numlist_take_Suc_at_term: assumes "i < length ns" shows "numlist ns ! (nllength (take (Suc i) ns) - 1) = \" proof - have "nllength (take (Suc i) ns) - 1 = nllength (take i ns) + nlength (ns ! i)" using nllength_take_Suc assms by (metis add_Suc_right diff_Suc_1) then show ?thesis using numlist_take_at_term assms by simp qed lemma nllength_take: assumes "i < length ns" shows "nllength (take i ns) + nlength (ns ! i) < nllength ns" proof - have "ns = take i ns @ [ns ! i] @ drop (Suc i) ns" using assms by (metis Cons_eq_appendI append_self_conv2 id_take_nth_drop) then have "numlist ns = numlist (take i ns) @ numlist [ns ! i] @ numlist (drop (Suc i) ns)" using numlist_append by metis then have "nllength ns = nllength (take i ns) + nllength [ns ! i] + nllength (drop (Suc i) ns)" by (simp add: nllength_def) then have "nllength ns \ nllength (take i ns) + nllength [ns ! i]" by simp then have "nllength ns \ nllength (take i ns) + Suc (nlength (ns ! i))" using nllength by simp then show ?thesis by simp qed text \ The contents of a tape starting with the start symbol @{text \} followed by the symbol sequence representing a list of numbers: \ definition nlcontents :: "nat list \ (nat \ symbol)" ("\_\\<^sub>N\<^sub>L") where "\ns\\<^sub>N\<^sub>L \ \numlist ns\" lemma clean_tape_nlcontents: "clean_tape (\ns\\<^sub>N\<^sub>L, i)" by (simp add: nlcontents_def proper_symbols_numlist) lemma nlcontents_Nil: "\[]\\<^sub>N\<^sub>L = \[]\" using nlcontents_def by (simp add: numlist_Nil) lemma nlcontents_rneigh_4: assumes "i < length ns" shows "rneigh (\ns\\<^sub>N\<^sub>L, Suc (nllength (take i ns))) {\} (nlength (ns ! i))" proof (rule rneighI) let ?tp = "(\ns\\<^sub>N\<^sub>L, Suc (nllength (take i ns)))" show "fst ?tp (snd ?tp + nlength (ns ! i)) \ {\}" proof - have "snd ?tp + nlength (ns ! i) \ nllength ns" using nllength_take assms by (simp add: Suc_leI) then have "fst ?tp (snd ?tp + nlength (ns ! i)) = numlist ns ! (nllength (take i ns) + nlength (ns ! i))" using nlcontents_def contents_inbounds nllength_def by simp then show ?thesis using numlist_take_at_term assms by simp qed show "fst ?tp (snd ?tp + j) \ {\}" if "j < nlength (ns ! i)" for j proof - have "snd ?tp + nlength (ns ! i) \ nllength ns" using nllength_take assms by (simp add: Suc_leI) then have "snd ?tp + j \ nllength ns" using that by simp then have "fst ?tp (snd ?tp + j) = numlist ns ! (nllength (take i ns) + j)" using nlcontents_def contents_inbounds nllength_def by simp then have "fst ?tp (snd ?tp + j) = canrepr (ns ! i) ! j" using assms that numlist_take_at by simp then show ?thesis using bit_symbols_canrepr that by fastforce qed qed lemma nlcontents_rneigh_04: assumes "i < length ns" shows "rneigh (\ns\\<^sub>N\<^sub>L, Suc (nllength (take i ns))) {\, \} (nlength (ns ! i))" proof (rule rneighI) let ?tp = "(\ns\\<^sub>N\<^sub>L, Suc (nllength (take i ns)))" show "fst ?tp (snd ?tp + nlength (ns ! i)) \ {\, \}" proof - have "snd ?tp + nlength (ns ! i) \ nllength ns" using nllength_take assms by (simp add: Suc_leI) then have "fst ?tp (snd ?tp + nlength (ns ! i)) = numlist ns ! (nllength (take i ns) + nlength (ns ! i))" using nlcontents_def contents_inbounds nllength_def by simp then show ?thesis using numlist_take_at_term assms by simp qed show "fst ?tp (snd ?tp + j) \ {\, \}" if "j < nlength (ns ! i)" for j proof - have "snd ?tp + nlength (ns ! i) \ nllength ns" using nllength_take assms by (simp add: Suc_leI) then have "snd ?tp + j \ nllength ns" using that by simp then have "fst ?tp (snd ?tp + j) = numlist ns ! (nllength (take i ns) + j)" using nlcontents_def contents_inbounds nllength_def by simp then have "fst ?tp (snd ?tp + j) = canrepr (ns ! i) ! j" using assms that numlist_take_at by simp then show ?thesis using bit_symbols_canrepr that by fastforce qed qed text \ A tape storing a list of numbers, with the tape head in the first blank cell after the symbols: \ abbreviation nltape :: "nat list \ tape" where "nltape ns \ (\ns\\<^sub>N\<^sub>L, Suc (nllength ns))" text \ A tape storing a list of numbers, with the tape head on the first symbol representing the $i$-th number, unless the $i$-th number is zero, in which case the tape head is on the terminator symbol of this zero. If $i$ is out of bounds of the list, the tape head is at the first blank after the list. \ abbreviation nltape' :: "nat list \ nat \ tape" where "nltape' ns i \ (\ns\\<^sub>N\<^sub>L, Suc (nllength (take i ns)))" lemma nltape'_tape_read: "|.| (nltape' ns i) = \ \ i \ length ns" proof - have "|.| (nltape' ns i) = \" if "i \ length ns" for i proof - have "nltape' ns i \ (\ns\\<^sub>N\<^sub>L, Suc (nllength ns))" using that by simp then show ?thesis using nlcontents_def contents_outofbounds nllength_def by (metis Suc_eq_plus1 add.left_neutral fst_conv less_Suc0 less_add_eq_less snd_conv) qed moreover have "|.| (nltape' ns i) \ \" if "i < length ns" for i using that nlcontents_def contents_inbounds nllength_def nllength_take proper_symbols_numlist by (metis Suc_leI add_Suc_right diff_Suc_1 fst_conv less_add_same_cancel1 less_le_trans not_add_less2 plus_1_eq_Suc snd_conv zero_less_Suc) ultimately show ?thesis by (meson le_less_linear) qed subsection \Moving to the next element\ text \ The next Turing machine provides a means to iterate over a list of numbers. If the TM starts in a configuration where the tape $j_1$ contains a list of numbers and the tape head is on the first symbol of the $i$-th element of this list, then after the TM has finished, the $i$-th element will be written on tape $j_2$ and the tape head on $j_1$ will have advanced by one list element. If $i$ is the last element of the list, the tape head on $j_1$ will be on a blank symbol. One can execute this TM in a loop until the tape head reaches a blank. The TM is generic over a parameter $z$ representing the terminator symbol, so it can be used for other kinds of lists, too (see Section~\ref{s:tm-numlistlist}). \null \ definition tm_nextract :: "symbol \ tapeidx \ tapeidx \ machine" where "tm_nextract z j1 j2 \ tm_erase_cr j2 ;; tm_cp_until j1 j2 {z} ;; tm_cr j2 ;; tm_right j1" lemma tm_nextract_tm: assumes "G \ 4" and "G > z" and "0 < j2" and "j2 < k" and "j1 < k" and "k \ 2" shows "turing_machine k G (tm_nextract z j1 j2)" unfolding tm_nextract_def using assms tm_erase_cr_tm tm_cp_until_tm tm_cr_tm tm_right_tm by simp text \ The next locale is for the case @{text "z = \"}. \null \ locale turing_machine_nextract_4 = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_erase_cr j2" definition "tm2 \ tm1 ;; tm_cp_until j1 j2 {\}" definition "tm3 \ tm2 ;; tm_cr j2" definition "tm4 \ tm3 ;; tm_right j1" lemma tm4_eq_tm_nextract: "tm4 = tm_nextract \ j1 j2" unfolding tm1_def tm2_def tm3_def tm4_def tm_nextract_def by simp context fixes tps0 :: "tape list" and k idx dummy :: nat and ns :: "nat list" assumes jk: "j1 < k" "j2 < k" "0 < j1" "0 < j2" "j1 \ j2" "length tps0 = k" and idx: "idx < length ns" and tps0: "tps0 ! j1 = nltape' ns idx" "tps0 ! j2 = (\dummy\\<^sub>N, 1)" begin definition "tps1 \ tps0[j2 := (\0\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 7 + 2 * nlength dummy" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: jk idx tps0 tps1_def assms) show "proper_symbols (canrepr dummy)" using proper_symbols_canrepr by simp show "tps1 = tps0[j2 := (\[]\, 1)]" using ncontents_0 tps1_def by simp qed definition "tps2 \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, nllength (take (Suc idx) ns)), j2 := (\ns ! idx\\<^sub>N, Suc (nlength (ns ! idx)))]" lemma tm2 [transforms_intros]: assumes "ttt = 7 + 2 * nlength dummy + Suc (nlength (ns ! idx))" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: jk idx tps0 tps2_def tps1_def) show "rneigh (tps1 ! j1) {\} (nlength (ns ! idx))" using tps1_def tps0 jk by (simp add: idx nlcontents_rneigh_4) show "tps2 = tps1 [j1 := tps1 ! j1 |+| nlength (ns ! idx), j2 := implant (tps1 ! j1) (tps1 ! j2) (nlength (ns ! idx))]" (is "?l = ?r") proof (rule nth_equalityI) show len: "length ?l = length ?r" using tps1_def tps2_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - consider "i = j1" | "i = j2" | "i \ j1 \ i \ j2" by auto then show ?thesis proof (cases) case 1 then show ?thesis using tps0 that tps1_def tps2_def jk nllength_take_Suc[OF idx] idx by simp next case 2 then have lhs: "?l ! i = (\ns ! idx\\<^sub>N, Suc (nlength (ns ! idx)))" using tps2_def jk by simp let ?i = "Suc (nllength (take idx ns))" have i1: "?i > 0" by simp have "nlength (ns ! idx) + (?i - 1) \ nllength ns" using idx by (simp add: add.commute less_or_eq_imp_le nllength_take) then have i2: "nlength (ns ! idx) + (?i - 1) \ length (numlist ns)" using nllength_def by simp have "?r ! i = implant (tps1 ! j1) (tps1 ! j2) (nlength (ns ! idx))" using 2 tps1_def jk by simp also have "... = implant (\ns\\<^sub>N\<^sub>L, ?i) (\0\\<^sub>N, 1) (nlength (ns ! idx))" using tps1_def jk tps0 by simp also have "... = (\[] @ (take (nlength (ns ! idx)) (drop (?i - 1) (numlist ns)))\, Suc (length []) + nlength (ns ! idx))" using implant_contents[OF i1 i2] by (metis One_nat_def list.size(3) ncontents_0 nlcontents_def) finally have "?r ! i = (\[] @ (take (nlength (ns ! idx)) (drop (?i - 1) (numlist ns)))\, Suc (length []) + nlength (ns ! idx))" . then have "?r ! i = (\take (nlength (ns ! idx)) (drop (nllength (take idx ns)) (numlist ns))\, Suc (nlength (ns ! idx)))" by simp then have "?r ! i = (\canrepr (ns ! idx)\, Suc (nlength (ns ! idx)))" using take_drop_numlist'[OF idx] by simp then show ?thesis using lhs by simp next case 3 then show ?thesis using that tps1_def tps2_def jk by simp qed qed qed show "ttt = 7 + 2 * nlength dummy + Suc (nlength (ns ! idx))" using assms(1) . qed definition "tps3 \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, nllength (take (Suc idx) ns)), j2 := (\ns ! idx\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 11 + 2 * nlength dummy + 2 * (nlength (ns ! idx))" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: jk idx tps0 tps2_def tps3_def time: assms tps2_def jk) show "clean_tape (tps2 ! j2)" using tps2_def jk clean_tape_ncontents by simp qed definition "tps4 \ tps0 [j1 := nltape' ns (Suc idx), j2 := (\ns ! idx\\<^sub>N, 1)]" lemma tm4: assumes "ttt = 12 + 2 * nlength dummy + 2 * (nlength (ns ! idx))" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: jk idx tps0 tps3_def tps4_def time: assms) show "tps4 = tps3[j1 := tps3 ! j1 |+| 1]" using tps3_def tps4_def jk tps0 by (metis Suc_eq_plus1 fst_conv list_update_overwrite list_update_swap nth_list_update_eq nth_list_update_neq snd_conv) qed end (* context *) end (* locale turing_machine_nextract *) lemma transforms_tm_nextract_4I [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k idx dummy :: nat and ns :: "nat list" assumes "j1 < k" "j2 < k" "0 < j1" "0 < j2" "j1 \ j2" "length tps = k" and "idx < length ns" assumes "tps ! j1 = nltape' ns idx" "tps ! j2 = (\dummy\\<^sub>N, 1)" assumes "ttt = 12 + 2 * nlength dummy + 2 * (nlength (ns ! idx))" assumes "tps' = tps [j1 := nltape' ns (Suc idx), j2 := (\ns ! idx\\<^sub>N, 1)]" shows "transforms (tm_nextract \ j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_nextract_4 j1 j2 . show ?thesis using assms loc.tm4 loc.tps4_def loc.tm4_eq_tm_nextract by simp qed subsection \Appending an element\ text \ The next Turing machine appends the number on tape $j_2$ to the list of numbers on tape $j_1$. \ definition tm_append :: "tapeidx \ tapeidx \ machine" where "tm_append j1 j2 \ tm_right_until j1 {\} ;; tm_cp_until j2 j1 {\} ;; tm_cr j2 ;; tm_char j1 \" lemma tm_append_tm: assumes "0 < j1" and "G \ 5" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_append j1 j2)" unfolding tm_append_def using assms tm_right_until_tm tm_cp_until_tm tm_right_tm tm_char_tm tm_cr_tm by simp locale turing_machine_append = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_right_until j1 {\}" definition "tm2 \ tm1 ;; tm_cp_until j2 j1 {\}" definition "tm3 \ tm2 ;; tm_cr j2" definition "tm4 \ tm3 ;; tm_char j1 \" lemma tm4_eq_tm_append: "tm4 = tm_append j1 j2" unfolding tm4_def tm3_def tm2_def tm1_def tm_append_def by simp context fixes tps0 :: "tape list" and k i1 n :: nat and ns :: "nat list" assumes jk: "length tps0 = k" "j1 < k" "j2 < k" "j1 \ j2" "0 < j1" and i1: "i1 \ Suc (nllength ns)" and tps0: "tps0 ! j1 = (\ns\\<^sub>N\<^sub>L, i1)" "tps0 ! j2 = (\n\\<^sub>N, 1)" begin lemma k: "k \ 2" using jk by simp lemma tm1 [transforms_intros]: assumes "ttt = Suc (Suc (nllength ns) - i1)" and "tps' = tps0[j1 := nltape ns]" shows "transforms tm1 tps0 ttt tps'" unfolding tm1_def proof (tform tps: jk k) let ?l = "Suc (nllength ns) - i1" show "rneigh (tps0 ! j1) {0} ?l" proof (rule rneighI) show "(tps0 ::: j1) (tps0 :#: j1 + ?l) \ {0}" using tps0 jk nlcontents_def nllength_def by simp show "(tps0 ::: j1) (tps0 :#: j1 + i) \ {0}" if "i < Suc (nllength ns) - i1" for i proof - have "i1 + i < Suc (nllength ns)" using that i1 by simp then show ?thesis using proper_symbols_numlist nllength_def tps0 nlcontents_def contents_def by (metis One_nat_def Suc_le_eq diff_Suc_1 fst_eqD less_Suc_eq_0_disj less_nat_zero_code singletonD snd_eqD) qed qed show "ttt = Suc (Suc (nllength ns) - i1)" using assms(1) . show "tps' = tps0[j1 := tps0 ! j1 |+| Suc (nllength ns) - i1]" using assms(2) tps0 i1 by simp qed lemma tm2 [transforms_intros]: assumes "ttt = 3 + nllength ns - i1 + nlength n" and "tps' = tps0 [j1 := (\numlist ns @ canrepr n\, Suc (nllength ns) + nlength n), j2 := (\n\\<^sub>N, Suc (nlength n))]" shows "transforms tm2 tps0 ttt tps'" unfolding tm2_def proof (tform tps: jk tps0) let ?tps = "tps0[j1 := nltape ns]" have j1: "?tps ! j1 = nltape ns" by (simp add: jk) have j2: "?tps ! j2 = (\n\\<^sub>N, 1)" using tps0 jk by simp show "rneigh (tps0[j1 := nltape ns] ! j2) {0} (nlength n)" proof (rule rneighI) show "(?tps ::: j2) (?tps :#: j2 + nlength n) \ {0}" using j2 contents_outofbounds by simp show "\i. i < nlength n \ (?tps ::: j2) (?tps :#: j2 + i) \ {0}" using j2 tps0 bit_symbols_canrepr by fastforce qed show "tps' = tps0 [j1 := nltape ns, j2 := ?tps ! j2 |+| nlength n, j1 := implant (?tps ! j2) (?tps ! j1) (nlength n)]" (is "_ = ?rhs") proof - have "implant (?tps ! j2) (?tps ! j1) (nlength n) = implant (\n\\<^sub>N, 1) (nltape ns) (nlength n)" using j1 j2 by simp also have "... = (\numlist ns @ (take (nlength n) (drop (1 - 1) (canrepr n)))\, Suc (length (numlist ns)) + nlength n)" using implant_contents nlcontents_def nllength_def by simp also have "... = (\numlist ns @ canrepr n\, Suc (length (numlist ns)) + nlength n)" by simp also have "... = (\numlist ns @ canrepr n\, Suc (nllength ns) + nlength n)" using nllength_def by simp also have "... = tps' ! j1" by (metis assms(2) jk(1,2,4) nth_list_update_eq nth_list_update_neq) finally have "implant (?tps ! j2) (?tps ! j1) (nlength n) = tps' ! j1" . then have "tps' ! j1 = implant (?tps ! j2) (?tps ! j1) (nlength n)" by simp then have "tps' ! j1 = ?rhs ! j1" by (simp add: jk) moreover have "tps' ! j2 = ?rhs ! j2" using assms(2) jk j2 by simp moreover have "tps' ! j = ?rhs ! j" if "j < length tps'" "j \ j1" "j \ j2" for j using that assms(2) by simp moreover have "length tps' = length ?rhs" using assms(2) by simp ultimately show ?thesis using nth_equalityI by blast qed show "ttt = Suc (Suc (nllength ns) - i1) + Suc (nlength n)" using assms(1) j1 tps0 i1 by simp qed definition "tps3 \ tps0 [j1 := (\numlist ns @ canrepr n\, Suc (nllength ns) + nlength n)]" lemma tm3 [transforms_intros]: assumes "ttt = 6 + nllength ns - i1 + 2 * nlength n" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: jk k) let ?tp1 = "(\numlist ns @ canrepr n\, Suc (nllength ns) + nlength n)" let ?tp2 = "(\n\\<^sub>N, Suc (nlength n))" show "clean_tape (tps0 [j1 := ?tp1, j2 := ?tp2] ! j2)" by (simp add: clean_tape_ncontents jk) show "tps3 = tps0[j1 := ?tp1, j2 := ?tp2, j2 := tps0 [j1 := ?tp1, j2 := ?tp2] ! j2 |#=| 1]" using tps3_def tps0 jk by (metis (no_types, lifting) add_Suc fst_conv length_list_update list_update_id list_update_overwrite list_update_swap nth_list_update_eq) show "ttt = 3 + nllength ns - i1 + nlength n + (tps0[j1 := ?tp1, j2 := ?tp2] :#: j2 + 2)" proof - have "tps0[j1 := ?tp1, j2 := ?tp2] :#: j2 = Suc (nlength n)" by (simp add: jk) then show ?thesis using jk tps0 i1 assms(1) by simp qed qed definition "tps4 = tps0 [j1 := (\numlist (ns @ [n])\, Suc (nllength (ns @ [n])))]" lemma tm4: assumes "ttt = 7 + nllength ns - i1 + 2 * nlength n" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: jk tps0 tps3_def) show "ttt = 6 + nllength ns - i1 + 2 * nlength n + 1 " using i1 assms(1) by simp show "tps4 = tps3[j1 := tps3 ! j1 |:=| \ |+| 1]" (is "tps4 = ?tps") proof - have "?tps = tps0[j1 := (\numlist ns @ canrepr n\, Suc (nllength ns + nlength n)) |:=| \ |+| 1]" using tps3_def by (simp add: jk) moreover have "(\numlist ns @ canrepr n\, Suc (nllength ns + nlength n)) |:=| \ |+| 1 = (\numlist (ns @ [n])\, Suc (nllength (ns @ [n])))" (is "?lhs = ?rhs") proof - have "?lhs = (\numlist ns @ canrepr n\(Suc (nllength ns + nlength n) := \), Suc (Suc (nllength ns + nlength n)))" by simp also have "... = (\numlist ns @ canrepr n\(Suc (nllength ns + nlength n) := \), Suc (nllength (ns @ [n])))" using nllength_def numlist_def by simp also have "... = (\(numlist ns @ canrepr n) @ [\]\, Suc (nllength (ns @ [n])))" using contents_snoc by (metis length_append nllength_def) also have "... = (\numlist ns @ canrepr n @ [\]\, Suc (nllength (ns @ [n])))" by simp also have "... = (\numlist (ns @ [n])\, Suc (nllength (ns @ [n])))" using numlist_def by simp finally show "?lhs = ?rhs" . qed ultimately show ?thesis using tps4_def by auto qed qed end end (* locale turing_machine_append *) lemma tm_append [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps :: "tape list" and k i1 n :: nat and ns :: "nat list" assumes "0 < j1" assumes "length tps = k" "j1 < k" "j2 < k" "j1 \ j2" "i1 \ Suc (nllength ns)" and "tps ! j1 = (\ns\\<^sub>N\<^sub>L, i1)" and "tps ! j2 = (\n\\<^sub>N, 1)" assumes "ttt = 7 + nllength ns - i1 + 2 * nlength n" and "tps' = tps [j1 := nltape (ns @ [n])]" shows "transforms (tm_append j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_append j1 j2 . show ?thesis using loc.tm4 loc.tm4_eq_tm_append loc.tps4_def assms nlcontents_def by simp qed subsection \Computing the length\ text \ The next Turing machine counts the number of terminator symbols $z$ on tape $j_1$ and stores the result on tape $j_2$. Thus, if $j_1$ contains a list of numbers, tape $j_2$ will contain the length of the list. \ definition tm_count :: "tapeidx \ tapeidx \ symbol \ machine" where "tm_count j1 j2 z \ WHILE tm_right_until j1 {\, z} ; \rs. rs ! j1 \ \ DO tm_incr j2 ;; tm_right j1 DONE ;; tm_cr j1" lemma tm_count_tm: assumes "0 < j2" and "j1 < k" and "j2 < k" and "j1 \ j2" and "G \ 4" shows "turing_machine k G (tm_count j1 j2 z)" unfolding tm_count_def using turing_machine_loop_turing_machine tm_right_until_tm tm_incr_tm tm_cr_tm tm_right_tm assms by simp locale turing_machine_count = fixes j1 j2 :: tapeidx begin definition "tmC \ tm_right_until j1 {\, \}" definition "tmB1 \ tm_incr j2" definition "tmB2 \ tmB1 ;; tm_right j1" definition "tmL \ WHILE tmC ; \rs. rs ! j1 \ \ DO tmB2 DONE" definition "tm2 \ tmL ;; tm_cr j1" lemma tm2_eq_tm_count: "tm2 = tm_count j1 j2 \" unfolding tmB1_def tmB2_def tmC_def tmL_def tm2_def tm_count_def by simp context fixes tps0 :: "tape list" and k :: nat and ns :: "nat list" assumes jk: "j1 < k" "j2 < k" "0 < j2" "j1 \ j2" "length tps0 = k" and tps0: "tps0 ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" "tps0 ! j2 = (\0\\<^sub>N, 1)" begin definition "tpsL t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, Suc (nllength (take t ns))), j2 := (\t\\<^sub>N, 1)]" definition "tpsC t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, if t < length ns then nllength (take (Suc t) ns) else Suc (nllength ns)), j2 := (\t\\<^sub>N, 1)]" lemma tmC: assumes "t \ length ns" and "ttt = Suc (if t = length ns then 0 else nlength (ns ! t))" shows "transforms tmC (tpsL t) ttt (tpsC t)" unfolding tmC_def proof (tform tps: jk tpsL_def tpsC_def) let ?n = "if t = length ns then 0 else nlength (ns ! t)" have *: "tpsL t ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength (take t ns)))" using tpsL_def jk by simp show "rneigh (tpsL t ! j1) {\, \} ?n" proof (cases "t = length ns") case True then have "tpsL t ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength ns))" (is "_ = ?tp") using * by simp moreover from this have "fst ?tp (snd ?tp) \ {\, \}" by (simp add: nlcontents_def nllength_def) moreover have "?n = 0" using True by simp ultimately show ?thesis using rneighI by simp next case False then have "tpsL t ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength (take t ns)))" using * by simp moreover have "?n = nlength (ns ! t)" using False by simp ultimately show ?thesis using nlcontents_rneigh_04 by (simp add: False assms(1) le_neq_implies_less) qed show "tpsC t = (tpsL t) [j1 := tpsL t ! j1 |+| (if t = length ns then 0 else nlength (ns ! t))]" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using tpsC_def tpsL_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - consider "i = j1" | "i = j2" | "i \ j1 \ i \ j2" by auto then show ?thesis proof (cases) case 1 show ?thesis proof (cases "t = length ns") case True then show ?thesis using 1 by (simp add: jk(2,4) tpsC_def tpsL_def) next case False then have "t < length ns" using assms(1) by simp then show ?thesis using 1 nllength_take_Suc jk tpsC_def tpsL_def by simp qed next case 2 then show ?thesis by (simp add: jk(2,4,5) tpsC_def tpsL_def) next case 3 then show ?thesis by (simp add: jk(2,4) tpsC_def tpsL_def) qed qed qed show "ttt = Suc (if t = length ns then 0 else nlength (ns ! t))" using assms(2) . qed lemma tmC' [transforms_intros]: assumes "t \ length ns" shows "transforms tmC (tpsL t) (Suc (nllength ns)) (tpsC t)" proof - have "Suc (if t = length ns then 0 else nlength (ns ! t)) \ Suc (if t = length ns then 0 else nllength ns)" using assms member_le_nllength by simp then have "Suc (if t = length ns then 0 else nlength (ns ! t)) \ Suc (nllength ns)" by auto then show ?thesis using tmC transforms_monotone assms by metis qed definition "tpsB1 t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, nllength (take (Suc t) ns)), j2 := (\Suc t\\<^sub>N, 1)]" lemma tmB1 [transforms_intros]: assumes "t < length ns" and "ttt = 5 + 2 * nlength t" shows "transforms tmB1 (tpsC t) ttt (tpsB1 t)" unfolding tmB1_def by (tform tps: jk tpsC_def tpsB1_def assms) definition "tpsB2 t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, Suc (nllength (take (Suc t) ns))), j2 := (\Suc t\\<^sub>N, 1)]" lemma tmB2: assumes "t < length ns" and "ttt = 6 + 2 * nlength t" shows "transforms tmB2 (tpsC t) ttt (tpsB2 t)" unfolding tmB2_def by (tform tps: jk tpsB1_def tpsB2_def assms) lemma tpsB2_eq_tpsL: "tpsB2 t = tpsL (Suc t)" using tpsB2_def tpsL_def by simp lemma tmB2' [transforms_intros]: assumes "t < length ns" shows "transforms tmB2 (tpsC t) (6 + 2 * nllength ns) (tpsL (Suc t))" proof - have "nlength t \ nllength ns" using assms(1) length_le_nllength nlength_le_n by (meson le_trans less_or_eq_imp_le) then have "6 + 2 * nlength t \ 6 + 2 * nllength ns" by simp then show ?thesis using assms tmB2 transforms_monotone tpsB2_eq_tpsL by metis qed lemma tmL [transforms_intros]: assumes "ttt = 13 * nllength ns ^ 2 + 2" shows "transforms tmL (tpsL 0) ttt (tpsC (length ns))" unfolding tmL_def proof (tform) show "read (tpsC t) ! j1 \ \" if "t < length ns" for t proof - have "tpsC t ! j1 = (\ns\\<^sub>N\<^sub>L, if t < length ns then nllength (take (Suc t) ns) else Suc (nllength ns))" using tpsC_def jk by simp then have *: "tpsC t ! j1 = (\ns\\<^sub>N\<^sub>L, nllength (take (Suc t) ns))" (is "_ = ?tp") using that by simp have "fst ?tp (snd ?tp) = \ns\\<^sub>N\<^sub>L (nllength (take (Suc t) ns))" by simp also have "... = \numlist ns\ (nllength (take (Suc t) ns))" using nlcontents_def by simp also have "... = numlist ns ! (nllength (take (Suc t) ns) - 1)" using nllength that contents_inbounds nllength_def nllength_take nllength_take_Suc by (metis Suc_leI add_Suc_right less_nat_zero_code not_less_eq) also have "... = 4" using numlist_take_Suc_at_term[OF that] by simp finally have "fst ?tp (snd ?tp) = \" . then have "fst ?tp (snd ?tp) \ \" by simp then show ?thesis using * jk(1,5) length_list_update tapes_at_read' tpsC_def by metis qed show "\ read (tpsC (length ns)) ! j1 \ \" proof - have "tpsC (length ns) ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength ns))" (is "_ = ?tp") using tpsC_def jk by simp moreover have "fst ?tp (snd ?tp) = 0" by (simp add: nlcontents_def nllength_def) ultimately have "read (tpsC (length ns)) ! j1 = \" using jk(1,5) length_list_update tapes_at_read' tpsC_def by metis then show ?thesis by simp qed show "length ns * (Suc (nllength ns) + (6 + 2 * nllength ns) + 2) + Suc (nllength ns) + 1 \ ttt" proof - have "length ns * (Suc (nllength ns) + (6 + 2 * nllength ns) + 2) + Suc (nllength ns) + 1 = length ns * (9 + 3 * nllength ns) + nllength ns + 2" by simp also have "... \ nllength ns * (9 + 3 * nllength ns) + nllength ns + 2" by (simp add: length_le_nllength) also have "... = nllength ns * (10 + 3 * nllength ns) + 2" by algebra also have "... = 10 * nllength ns + 3 * nllength ns ^ 2 + 2" by algebra also have "... \ 10 * nllength ns ^ 2 + 3 * nllength ns ^ 2 + 2" by (meson add_mono_thms_linordered_semiring(1) le_eq_less_or_eq mult_le_mono2 power2_nat_le_imp_le) also have "... = 13 * nllength ns ^ 2 + 2" by simp finally show ?thesis using assms by simp qed qed definition "tps2 \ tps0 [j2 := (\length ns\\<^sub>N, 1)]" lemma tm2: assumes "ttt = 13 * nllength ns ^ 2 + 5 + nllength ns" shows "transforms tm2 (tpsL 0) ttt tps2" unfolding tm2_def proof (tform tps: jk tpsC_def tps2_def) have *: "tpsC (length ns) ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength ns))" using jk tpsC_def by simp then show "clean_tape (tpsC (length ns) ! j1)" by (simp add: clean_tape_nlcontents) show "tps2 = (tpsC (length ns))[j1 := tpsC (length ns) ! j1 |#=| 1]" using jk tps0(1) tps2_def tpsC_def * by (metis fstI list_update_id list_update_overwrite list_update_swap) show "ttt = 13 * (nllength ns)\<^sup>2 + 2 + (tpsC (length ns) :#: j1 + 2)" using assms * by simp qed lemma tpsL_eq_tps0: "tpsL 0 = tps0" using tpsL_def tps0 by (metis One_nat_def list_update_id nllength_Nil take0) lemma tm2': assumes "ttt = 14 * nllength ns ^ 2 + 5" shows "transforms tm2 tps0 ttt tps2" proof - have "nllength ns \ nllength ns ^ 2" using power2_nat_le_imp_le by simp then have "13 * nllength ns ^ 2 + 5 + nllength ns \ ttt" using assms by simp then show ?thesis using assms tm2 transforms_monotone tpsL_eq_tps0 by simp qed end (* context tps0 *) end (* locale *) lemma transforms_tm_count_4I [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k :: nat and ns :: "nat list" assumes "j1 < k" "j2 < k" "0 < j2" "j1 \ j2" "length tps = k" assumes "tps ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" "tps ! j2 = (\0\\<^sub>N, 1)" assumes "ttt = 14 * nllength ns ^ 2 + 5" assumes "tps' = tps[j2 := (\length ns\\<^sub>N, 1)]" shows "transforms (tm_count j1 j2 \) tps ttt tps'" proof - interpret loc: turing_machine_count j1 j2 . show ?thesis using loc.tm2_eq_tm_count loc.tm2' loc.tps2_def assms by simp qed subsection \Extracting the $n$-th element\ text \ The next Turing machine expects a list on tape $j_1$ and an index $i$ on $j_2$ and writes the $i$-th element of the list to $j_2$, overwriting $i$. The TM does not terminate for out-of-bounds access, which of course we will never attempt. Again the parameter $z$ is a generic terminator symbol. \ definition tm_nth_inplace :: "tapeidx \ tapeidx \ symbol \ machine" where "tm_nth_inplace j1 j2 z \ WHILE [] ; \rs. rs ! j2 \ \ DO tm_decr j2 ;; tm_right_until j1 {z} ;; tm_right j1 DONE ;; tm_cp_until j1 j2 {z} ;; tm_cr j2 ;; tm_cr j1" lemma tm_nth_inplace_tm: assumes "k \ 2" and "G \ 4" and "0 < j2" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_nth_inplace j1 j2 \)" unfolding tm_nth_inplace_def using assms tm_decr_tm tm_right_until_tm tm_right_tm tm_cp_until_tm tm_cr_tm Nil_tm by (simp add: assms(1) turing_machine_loop_turing_machine) locale turing_machine_nth_inplace = fixes j1 j2 :: tapeidx begin definition "tmL1 \ tm_decr j2" definition "tmL2 \ tmL1 ;; tm_right_until j1 {\}" definition "tmL3 \ tmL2 ;; tm_right j1" definition "tmL \ WHILE [] ; \rs. rs ! j2 \ \ DO tmL3 DONE" definition "tm2 \ tmL ;; tm_cp_until j1 j2 {\}" definition "tm3 \ tm2 ;; tm_cr j2" definition "tm4 \ tm3 ;; tm_cr j1" lemma tm4_eq_tm_nth_inplace: "tm4 = tm_nth_inplace j1 j2 \" unfolding tmL1_def tmL2_def tmL3_def tmL_def tm2_def tm3_def tm4_def tm_nth_inplace_def by simp context fixes tps0 :: "tape list" and k idx :: nat and ns :: "nat list" assumes jk: "j1 < k" "j2 < k" "0 < j2" "j1 \ j2" "length tps0 = k" and idx: "idx < length ns" and tps0: "tps0 ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" "tps0 ! j2 = (\idx\\<^sub>N, 1)" begin definition "tpsL t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, Suc (nllength (take t ns))), j2 := (\idx - t\\<^sub>N, 1)]" definition "tpsL1 t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, Suc (nllength (take t ns))), j2 := (\idx - t - 1\\<^sub>N, 1)]" lemma tmL1 [transforms_intros]: assumes "ttt = 8 + 2 * nlength (idx - t)" shows "transforms tmL1 (tpsL t) ttt (tpsL1 t)" unfolding tmL1_def by (tform tps: tpsL_def tpsL1_def jk time: assms) definition "tpsL2 t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, nllength (take (Suc t) ns)), j2 := (\idx - t - 1\\<^sub>N, 1)]" lemma tmL2: assumes "t < length ns" and "ttt = 8 + 2 * nlength (idx - t) + Suc (nlength (ns ! t))" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" unfolding tmL2_def proof (tform tps: jk tpsL1_def tpsL2_def time: assms(2)) let ?l = "nlength (ns ! t)" show "rneigh (tpsL1 t ! j1) {\} ?l" proof - have "tpsL1 t ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength (take t ns)))" using tpsL1_def jk by (simp only: nth_list_update_eq nth_list_update_neq) then show ?thesis using assms(1) nlcontents_rneigh_4 by simp qed show "tpsL2 t = (tpsL1 t)[j1 := tpsL1 t ! j1 |+| nlength (ns ! t)]" (is "?l = ?r") proof (rule nth_equalityI) show len: "length ?l = length ?r" using tpsL1_def tpsL2_def jk by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - consider "i = j1" | "i = j2" | "i \ j1 \ i \ j2" by auto then show ?thesis proof (cases) case 1 then show ?thesis using that tpsL1_def tpsL2_def jk nllength_take_Suc[OF assms(1)] by simp next case 2 then show ?thesis using that tpsL1_def tpsL2_def jk by (simp only: nth_list_update_eq nth_list_update_neq' length_list_update) next case 3 then show ?thesis using that tpsL1_def tpsL2_def jk by (simp only: nth_list_update_eq jk nth_list_update_neq' length_list_update) qed qed qed qed lemma tmL2' [transforms_intros]: assumes "t < length ns" and "ttt = 9 + 2 * nlength idx + nlength (Max (set ns))" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" proof - let ?ttt = "8 + 2 * nlength (idx - t) + Suc (nlength (ns ! t))" have "transforms tmL2 (tpsL t) ?ttt (tpsL2 t)" using tmL2 assms by simp moreover have "ttt \ ?ttt" proof - have "nlength (idx - t) \ nlength idx" using nlength_mono by simp moreover have "nlength (ns ! t) \ nlength (Max (set ns))" using nlength_mono assms by simp ultimately show ?thesis using assms(2) by simp qed ultimately show ?thesis using transforms_monotone by simp qed definition "tpsL3 t \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, Suc (nllength (take (Suc t) ns))), j2 := (\idx - t - 1\\<^sub>N, 1)]" lemma tmL3: assumes "t < length ns" and "ttt = 10 + 2 * nlength idx + nlength (Max (set ns))" shows "transforms tmL3 (tpsL t) ttt (tpsL3 t)" unfolding tmL3_def proof (tform tps: jk tpsL2_def tpsL3_def assms(1) time: assms(2)) show "tpsL3 t = (tpsL2 t)[j1 := tpsL2 t ! j1 |+| 1]" using tpsL3_def tpsL2_def jk tps0 by (metis Groups.add_ac(2) fst_conv list_update_overwrite list_update_swap nth_list_update nth_list_update_neq plus_1_eq_Suc snd_conv) qed lemma tpsL3_eq_tpsL: "tpsL3 t = tpsL (Suc t)" using tpsL3_def tpsL_def by simp lemma tmL: assumes "ttt = idx * (12 + 2 * nlength idx + nlength (Max (set ns))) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL idx)" unfolding tmL_def proof (tform) let ?t = "10 + 2 * nlength idx + nlength (Max (set ns))" show "\t. t < idx \ transforms tmL3 (tpsL t) ?t (tpsL (Suc t))" using tmL3 tpsL3_eq_tpsL idx by simp show "read (tpsL t) ! j2 \ \" if "t < idx" for t proof - have "tpsL t ! j2 = (\idx - t\\<^sub>N, 1)" using tpsL_def jk by simp then have "read (tpsL t) ! j2 = \idx - t\\<^sub>N 1" using tapes_at_read' jk tpsL_def by (metis fst_conv length_list_update snd_conv) moreover have "idx - t > 0" using that by simp ultimately show "read (tpsL t) ! j2 \ \" using ncontents_1_blank_iff_zero by simp qed show "\ read (tpsL idx) ! j2 \ \" proof - have "tpsL idx ! j2 = (\idx - idx\\<^sub>N, 1)" using tpsL_def jk by simp then have "read (tpsL idx) ! j2 = \idx - idx\\<^sub>N 1" using tapes_at_read' jk tpsL_def by (metis fst_conv length_list_update snd_conv) then show ?thesis using ncontents_1_blank_iff_zero by simp qed show "idx * (10 + 2 * nlength idx + nlength (Max (set ns)) + 2) + 1 \ ttt" using assms by simp qed definition "tps1 \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, Suc (nllength (take idx ns))), j2 := (\0\\<^sub>N, 1)]" lemma tps1_eq_tpsL: "tps1 = tpsL idx" using tps1_def tpsL_def by simp lemma tps0_eq_tpsL: "tps0 = tpsL 0" using tps0 tpsL_def nllength_Nil by (metis One_nat_def list_update_id minus_nat.diff_0 take0) lemma tmL' [transforms_intros]: assumes "ttt = idx * (12 + 2 * nlength idx + nlength (Max (set ns))) + 1" shows "transforms tmL tps0 ttt tps1" using tmL assms tps0_eq_tpsL tps1_eq_tpsL by simp definition "tps2 \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, nllength (take (Suc idx) ns)), j2 := (\ns ! idx\\<^sub>N, Suc (nlength (ns ! idx)))]" lemma tm2 [transforms_intros]: assumes "ttt = idx * (12 + 2 * nlength idx + nlength (Max (set ns))) + 2 + nlength (ns ! idx)" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: jk tps2_def tps1_def time: assms) have "tps1 ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength (take idx ns)))" using tps1_def tps0 jk by simp then show "rneigh (tps1 ! j1) {\} (nlength (ns ! idx))" by (simp add: idx nlcontents_rneigh_4) show "tps2 = tps1 [j1 := tps1 ! j1 |+| nlength (ns ! idx), j2 := implant (tps1 ! j1) (tps1 ! j2) (nlength (ns ! idx))]" (is "?l = ?r") proof (rule nth_equalityI) show len: "length ?l = length ?r" using tps1_def tps2_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - consider "i = j1" | "i = j2" | "i \ j1 \ i \ j2" by auto then show ?thesis proof (cases) case 1 then show ?thesis using that tps1_def tps2_def jk nllength_take_Suc idx by simp next case 2 then have lhs: "?l ! i = (\ns ! idx\\<^sub>N, Suc (nlength (ns ! idx)))" using tps2_def jk by simp let ?i = "Suc (nllength (take idx ns))" have i1: "?i > 0" by simp have "nlength (ns ! idx) + (?i - 1) \ nllength ns" using idx by (simp add: add.commute less_or_eq_imp_le nllength_take) then have i2: "nlength (ns ! idx) + (?i - 1) \ length (numlist ns)" using nllength_def by simp have "?r ! i = implant (tps1 ! j1) (tps1 ! j2) (nlength (ns ! idx))" using 2 tps1_def jk by simp also have "... = implant (\ns\\<^sub>N\<^sub>L, ?i) (\0\\<^sub>N, 1) (nlength (ns ! idx))" proof - have "tps1 ! j1 = (\ns\\<^sub>N\<^sub>L, Suc (nllength (take idx ns)))" using tps1_def jk by simp moreover have "tps1 ! j2 = (\0\\<^sub>N, 1)" using tps1_def jk by simp ultimately show ?thesis by simp qed also have "... = (\[] @ (take (nlength (ns ! idx)) (drop (?i - 1) (numlist ns)))\, Suc (length []) + nlength (ns ! idx))" using implant_contents[OF i1 i2] by (metis One_nat_def list.size(3) ncontents_0 nlcontents_def) finally have "?r ! i = (\[] @ (take (nlength (ns ! idx)) (drop (?i - 1) (numlist ns)))\, Suc (length []) + nlength (ns ! idx))" . then have "?r ! i = (\take (nlength (ns ! idx)) (drop (nllength (take idx ns)) (numlist ns))\, Suc (nlength (ns ! idx)))" by simp then have "?r ! i = (\canrepr (ns ! idx)\, Suc (nlength (ns ! idx)))" using take_drop_numlist'[OF idx] by simp then show ?thesis using lhs by simp next case 3 then show ?thesis using that tps1_def tps2_def jk by simp qed qed qed qed definition "tps3 \ tps0 [j1 := (\ns\\<^sub>N\<^sub>L, nllength (take (Suc idx) ns)), j2 := (\ns ! idx\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = idx * (12 + 2 * nlength idx + nlength (Max (set ns))) + 5 + 2 * nlength (ns ! idx)" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def by (tform tps: clean_tape_ncontents jk tps2_def tps3_def time: assms tps2_def jk) definition "tps4 \ tps0 [j2 := (\ns ! idx\\<^sub>N, 1)]" lemma tm4: assumes "ttt = idx * (12 + 2 * nlength idx + nlength (Max (set ns))) + 7 + 2 * nlength (ns ! idx) + nllength (take (Suc idx) ns)" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: clean_tape_nlcontents jk tps3_def tps4_def time: assms jk tps3_def) show "tps4 = tps3[j1 := tps3 ! j1 |#=| 1]" using tps4_def tps3_def jk tps0(1) list_update_id[of tps0 j1] by (simp add: list_update_swap) qed lemma tm4': assumes "ttt = 18 * nllength ns ^ 2 + 12" shows "transforms tm4 tps0 ttt tps4" proof - let ?ttt = "idx * (12 + 2 * nlength idx + nlength (Max (set ns))) + 7 + 2 * nlength (ns ! idx) + nllength (take (Suc idx) ns)" have 1: "idx \ nllength ns" using idx length_le_nllength by (meson le_trans less_or_eq_imp_le) then have 2: "nlength idx \ nllength ns" using nlength_mono nlength_le_n by (meson dual_order.trans) have "ns \ []" using idx by auto then have 3: "nlength (Max (set ns)) \ nllength ns" using member_le_nllength by simp have 4: "nlength (ns ! idx) \ nllength ns" using idx member_le_nllength by simp have 5: "nllength (take (Suc idx) ns) \ nllength ns" by (metis Suc_le_eq add_Suc_right idx nllength_take nllength_take_Suc) have "?ttt \ idx * (12 + 2 * nllength ns + nllength ns) + 7 + 2 * nllength ns + nllength ns" using 2 3 4 5 by (meson add_le_mono le_eq_less_or_eq mult_le_mono2) also have "... = idx * (12 + 3 * nllength ns) + 7 + 3 * nllength ns" by simp also have "... \ idx * (12 + 3 * nllength ns) + (12 + 3 * nllength ns)" by simp also have "... = Suc idx * (12 + 3 * nllength ns)" by simp also have "... \ Suc (nllength ns) * (12 + 3 * nllength ns)" using 1 by simp also have "... = nllength ns * (12 + 3 * nllength ns) + (12 + 3 * nllength ns)" by simp also have "... = 12 * nllength ns + 3 * nllength ns ^ 2 + 12 + 3 * nllength ns" by algebra also have "... = 15 * nllength ns + 3 * nllength ns ^ 2 + 12" by simp also have "... \ 18 * nllength ns ^ 2 + 12" by (simp add: power2_eq_square) finally have "?ttt \ 18 * nllength ns ^ 2 + 12" . then show ?thesis using tm4 transforms_monotone assms by simp qed end (* context tps0 *) end (* locale turing_machine_nth_inplace *) lemma transforms_tm_nth_inplace_4I [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k idx :: nat and ns :: "nat list" assumes "j1 < k" "j2 < k" "0 < j2" "j1 \ j2" "length tps = k" and "idx < length ns" assumes "tps ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" "tps ! j2 = (\idx\\<^sub>N, 1)" assumes "ttt = 18 * nllength ns ^ 2 + 12" assumes "tps' = tps [j2 := (\ns ! idx\\<^sub>N, 1)]" shows "transforms (tm_nth_inplace j1 j2 \) tps ttt tps'" proof - interpret loc: turing_machine_nth_inplace j1 j2 . show ?thesis using assms loc.tm4_eq_tm_nth_inplace loc.tm4' loc.tps4_def by simp qed text \ The next Turing machine expects a list on tape $j_1$ and an index $i$ on tape $j_2$. It writes the $i$-th element of the list to tape $j_3$. Like the previous TM, it will not terminate on out-of-bounds access, and $z$ is a parameter for the symbol that terminates the list elements. \ definition tm_nth :: "tapeidx \ tapeidx \ tapeidx \ symbol \ machine" where "tm_nth j1 j2 j3 z \ tm_copyn j2 j3 ;; tm_nth_inplace j1 j3 z" lemma tm_nth_tm: assumes "k \ 2" and "G \ 4" and "0 < j2" "0 < j1" "j1 < k" "j2 < k" "0 < j3" "j3 < k" "j2 \ j3" shows "turing_machine k G (tm_nth j1 j2 j3 \)" unfolding tm_nth_def using tm_copyn_tm tm_nth_inplace_tm assms by simp lemma transforms_tm_nth_4I [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and k idx :: nat and ns :: "nat list" assumes "j1 < k" "j2 < k" "j3 < k" "0 < j1" "0 < j2" "0 < j3" "j1 \ j2" "j2 \ j3" "j1 \ j3" and "length tps = k" and "idx < length ns" assumes "tps ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" "tps ! j2 = (\idx\\<^sub>N, 1)" "tps ! j3 = (\0\\<^sub>N, 1)" assumes "ttt = 21 * nllength ns ^ 2 + 26" assumes "tps' = tps [j3 := (\ns ! idx\\<^sub>N, 1)]" shows "transforms (tm_nth j1 j2 j3 \) tps ttt tps'" proof - let ?ttt = "14 + 3 * (nlength idx + nlength 0) + (18 * (nllength ns)\<^sup>2 + 12)" have "transforms (tm_nth j1 j2 j3 \) tps ?ttt tps'" unfolding tm_nth_def proof (tform tps: assms(1-11)) show "tps ! j2 = (\idx\\<^sub>N, 1)" using assms by simp show "tps ! j3 = (\0\\<^sub>N, 1)" using assms by simp show "tps[j3 := (\idx\\<^sub>N, 1)] ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" using assms by simp then show "tps' = tps [j3 := (\idx\\<^sub>N, 1), j3 := (\ns ! idx\\<^sub>N, 1)]" using assms by (metis One_nat_def list_update_overwrite) qed moreover have "?ttt \ ttt" proof - have "nlength idx \ idx" using nlength_le_n by simp then have "nlength idx \ length ns" using assms(11) by simp then have "nlength idx \ nllength ns" using length_le_nllength by (meson order.trans) then have "nlength idx \ nllength ns ^ 2" by (meson le_refl order_trans power2_nat_le_imp_le) moreover have "?ttt = 3 * nlength idx + 18 * (nllength ns)\<^sup>2 + 26" by simp ultimately show ?thesis using assms(15) by simp qed ultimately show ?thesis using transforms_monotone by simp qed subsection \Finding the previous position of an element\ text \ The Turing machine in this section implements a slightly peculiar functionality, which we will start using only in Section~\ref{s:Reducing}. Given a list of numbers and an index $i$ it determines the greatest index less than $i$ where the list contains the same element as in position $i$. If no such element exists, it returns $i$. Formally: \ definition previous :: "nat list \ nat \ nat" where "previous ns idx \ if \i ns ! i = ns ! idx else idx" lemma previous_less: assumes "\i ns ! (previous ns idx) = ns ! idx" proof - let ?P = "\i. i < idx \ ns ! i = ns ! idx" have "previous ns idx = (GREATEST i. ?P i)" using assms previous_def by simp moreover have "\y. ?P y \ y \ idx" by simp ultimately show ?thesis using GreatestI_ex_nat[OF assms, of idx] by simp qed lemma previous_eq: "previous ns idx = idx \ \ (\i idx" using previous_eq previous_less by (metis less_or_eq_imp_le) text \ The following Turing machine expects the list on tape $j_1$ and the index on tape $j_2$. It outputs the result on tape $j_2 + 5$. The tapes $j_2 + 1, \dots, j_2 + 4$ are scratch space. \ definition tm_prev :: "tapeidx \ tapeidx \ machine" where "tm_prev j1 j2 \ tm_copyn j2 (j2 + 5) ;; tm_nth j1 j2 (j2 + 1) \ ;; WHILE tm_equalsn (j2 + 2) j2 (j2 + 4) ; \rs. rs ! (j2 + 4) = \ DO tm_nth j1 (j2 + 2) (j2 + 3) \ ;; tm_equalsn (j2 + 1) (j2 + 3) (j2 + 4) ;; tm_setn (j2 + 3) 0 ;; IF \rs. rs ! (j2 + 4) \ \ THEN tm_setn (j2 + 4) 0 ;; tm_copyn (j2 + 2) (j2 + 5) ELSE [] ENDIF ;; tm_incr (j2 + 2) DONE ;; tm_erase_cr (j2 + 1) ;; tm_erase_cr (j2 + 2) ;; tm_erase_cr (j2 + 4)" lemma tm_prev_tm: assumes "k \ j2 + 6" and "G \ 4" and "j1 < j2" and "0 < j1" shows "turing_machine k G (tm_prev j1 j2)" unfolding tm_prev_def using assms tm_copyn_tm tm_nth_tm tm_equalsn_tm tm_setn_tm tm_incr_tm Nil_tm tm_erase_cr_tm turing_machine_loop_turing_machine turing_machine_branch_turing_machine by simp locale turing_machine_prev = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_copyn j2 (j2 + 5)" definition "tm2 \ tm1 ;; tm_nth j1 j2 (j2 + 1) \" definition "tmC \ tm_equalsn (j2 + 2) j2 (j2 + 4)" definition "tmB1 \ tm_nth j1 (j2 + 2) (j2 + 3) \" definition "tmB2 \ tmB1 ;; tm_equalsn (j2 + 1) (j2 + 3) (j2 + 4)" definition "tmB3 \ tmB2 ;; tm_setn (j2 + 3) 0" definition "tmI1 \ tm_setn (j2 + 4) 0" definition "tmI2 \ tmI1 ;; tm_copyn (j2 + 2) (j2 + 5)" definition "tmI \ IF \rs. rs ! (j2 + 4) \ \ THEN tmI2 ELSE [] ENDIF" definition "tmB4 \ tmB3 ;; tmI" definition "tmB5 \ tmB4 ;; tm_incr (j2 + 2)" definition "tmL \ WHILE tmC ; \rs. rs ! (j2 + 4) = \ DO tmB5 DONE" definition "tm3 \ tm2 ;; tmL" definition "tm4 \ tm3 ;; tm_erase_cr (j2 + 1)" definition "tm5 \ tm4 ;; tm_erase_cr (j2 + 2)" definition "tm6 \ tm5 ;; tm_erase_cr (j2 + 4)" lemma tm6_eq_tm_prev: "tm6 = tm_prev j1 j2" unfolding tm_prev_def tm3_def tmL_def tmB5_def tmB4_def tmI_def tmI2_def tmI1_def tmB3_def tmB2_def tmB1_def tmC_def tm2_def tm1_def tm4_def tm5_def tm6_def by simp context fixes tps0 :: "tape list" and k idx :: nat and ns :: "nat list" assumes jk: "0 < j1" "j1 < j2" "j2 + 6 \ k" "length tps0 = k" and idx: "idx < length ns" and tps0: "tps0 ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" "tps0 ! j2 = (\idx\\<^sub>N, 1)" "tps0 ! (j2 + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j2 + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j2 + 3) = (\0\\<^sub>N, 1)" "tps0 ! (j2 + 4) = (\0\\<^sub>N, 1)" "tps0 ! (j2 + 5) = (\0\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j2 + 5 := (\idx\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 14 + 3 * nlength idx" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def by (tform tps: jk idx tps0 tps1_def assms nlength_0) definition "tps2 \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 5 := (\idx\\<^sub>N, 1)]" lemma tm2: assumes "ttt = 14 + 3 * nlength idx + (21 * (nllength ns)\<^sup>2 + 26)" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: jk idx tps0 tps1_def tps2_def time: assms) definition rv :: "nat \ nat" where "rv t \ if \i ns ! i = ns ! idx else idx" lemma rv_0: "rv 0 = idx" using rv_def by simp lemma rv_le: "rv t \ max t idx" proof (cases "\iky. y < t \ ns ! y = ns ! idx \ y \ t" by simp qed ultimately have "?Q (rv t)" by simp then have "rv t < t" by simp then show ?thesis by simp next case False then show ?thesis using rv_def by auto qed lemma rv_change: assumes "t < length ns" and "idx < length ns" and "ns ! t = ns ! idx" shows "rv (Suc t) = t" proof - let ?P = "\i. i < Suc t \ ns ! i = ns ! idx" have "rv (Suc t) = Greatest ?P" using assms(3) rv_def by auto moreover have "Greatest ?P = t" proof (rule Greatest_equality) show "t < Suc t \ ns ! t = ns ! idx" using assms(3) by simp show "\y. y < Suc t \ ns ! y = ns ! idx \ y \ t" using assms by simp qed ultimately show ?thesis by simp qed lemma rv_keep: assumes "t < length ns" and "idx < length ns" and "ns ! t \ ns ! idx" shows "rv (Suc t) = rv t" proof (cases "\iiky. y < t \ ns ! y = ns ! idx \ y \ t" by simp qed ultimately have "?Q (rv t)" by simp let ?P = "\i. i < Suc t \ ns ! i = ns ! idx" have "rv (Suc t) = Greatest ?P" using True rv_def by simp moreover have "Greatest ?P = rv t" proof (rule Greatest_equality) show "rv t < Suc t \ ns ! rv t = ns ! idx" using `?Q (rv t)` by simp show "y \ rv t" if "y < Suc t \ ns ! y = ns ! idx" for y proof - have "?Q y" using True assms(3) less_antisym that by blast moreover have "\y. ?Q y \ y \ t" by simp ultimately have "y \ Greatest ?Q" using Greatest_le_nat[of ?Q] by blast then show ?thesis using rvt by simp qed qed ultimately show ?thesis by simp next case False then show ?thesis using rv_def by auto qed definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 5 := (\rv t\\<^sub>N, 1)]" lemma tpsL_eq_tps2: "tpsL 0 = tps2" using tpsL_def tps2_def tps0 jk rv_0 by (metis add_eq_self_zero add_left_imp_eq gr_implies_not0 less_numeral_extra(1) list_update_id list_update_swap one_add_one) definition tpsC :: "nat \ tape list" where "tpsC t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 4 := (\t = idx\\<^sub>B, 1), j2 + 5 := (\rv t\\<^sub>N, 1)]" lemma tmC: assumes "ttt = 3 * nlength (min t idx) + 7" shows "transforms tmC (tpsL t) ttt (tpsC t)" unfolding tmC_def by (tform tps: jk idx tps0 tpsL_def tpsC_def time: assms) lemma tmC' [transforms_intros]: assumes "ttt = 3 * nllength ns ^ 2 + 7" and "t \ idx" shows "transforms tmC (tpsL t) ttt (tpsC t)" proof - have "nlength (min t idx) \ nllength ns" using idx assms(2) by (metis le_trans length_le_nllength less_or_eq_imp_le min_absorb1 nlength_le_n) then have "nlength (min t idx) \ nllength ns ^ 2" by (metis le_square min.absorb2 min.coboundedI1 power2_eq_square) then have "3 * nlength (min t idx) + 7 \ 3 * nllength ns ^ 2 + 7" by simp then show ?thesis using tmC assms(1) transforms_monotone by blast qed lemma condition_tpsC: "(read (tpsC t)) ! (j2 + 4) \ \ \ t = idx" using tpsC_def read_ncontents_eq_0 jk by simp definition "tpsB1 t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 3 := (\ns ! t\\<^sub>N, 1), j2 + 4 := (\t = idx\\<^sub>B, 1), j2 + 5 := (\rv t\\<^sub>N, 1)]" lemma tmB1 [transforms_intros]: assumes "ttt = 21 * (nllength ns)\<^sup>2 + 26" and "t < idx" shows "transforms tmB1 (tpsC t) ttt (tpsB1 t)" unfolding tmB1_def by (tform tps: jk idx tps0 tpsC_def tpsB1_def time: assms idx) definition "tpsB2 t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 3 := (\ns ! t\\<^sub>N, 1), j2 + 4 := (\ns ! idx = ns ! t\\<^sub>B, 1), j2 + 5 := (\rv t\\<^sub>N, 1)]" lemma tmB2: assumes "ttt = 21 * (nllength ns)\<^sup>2 + 26 + (3 * nlength (min (ns ! idx) (ns ! t)) + 7)" and "t < idx" shows "transforms tmB2 (tpsC t) ttt (tpsB2 t)" unfolding tmB2_def proof (tform tps: tpsB1_def jk assms(2) time: assms(1)) show "tpsB2 t = (tpsB1 t)[j2 + 4 := (\ns ! idx = ns ! t\\<^sub>B, 1)]" unfolding tpsB2_def tpsB1_def by (simp add: list_update_swap[of "j2 + 4"]) qed lemma tmB2' [transforms_intros]: assumes "ttt = 24 * (nllength ns)\<^sup>2 + 33" and "t < idx" shows "transforms tmB2 (tpsC t) ttt (tpsB2 t)" proof - let ?ttt = "21 * (nllength ns)\<^sup>2 + 26 + (3 * nlength (min (ns ! idx) (ns ! t)) + 7)" have "?ttt = 21 * (nllength ns)\<^sup>2 + 33 + 3 * nlength (min (ns ! idx) (ns ! t))" by simp also have "... \ 21 * (nllength ns)\<^sup>2 + 33 + 3 * nllength ns" using nlength_min_le_nllength idx assms(2) by simp also have "... \ 21 * (nllength ns)\<^sup>2 + 33 + 3 * nllength ns ^ 2" by (simp add: power2_eq_square) also have "... = 24 * (nllength ns)\<^sup>2 + 33" by simp finally have "?ttt \ 24 * (nllength ns)\<^sup>2 + 33" . then show ?thesis using assms tmB2 transforms_monotone by blast qed definition "tpsB3 t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 4 := (\ns ! idx = ns ! t\\<^sub>B, 1), j2 + 5 := (\rv t\\<^sub>N, 1)]" lemma condition_tpsB3: "(read (tpsB3 t)) ! (j2 + 4) \ \ \ ns ! idx = ns ! t" using tpsB3_def read_ncontents_eq_0 jk by simp lemma tmB3 [transforms_intros]: assumes "ttt = 24 * (nllength ns)\<^sup>2 + 33 + (10 + 2 * nlength (ns ! t))" and "t < idx" shows "transforms tmB3 (tpsC t) ttt (tpsB3 t)" unfolding tmB3_def proof (tform tps: assms(2) tpsB2_def jk time: assms(1)) show "tpsB3 t = (tpsB2 t)[j2 + 3 := (\0\\<^sub>N, 1)]" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using tpsB2_def tpsB3_def tps0 by simp show "?l ! j = ?r ! j" if "j < length ?l" for j proof - consider "j = j1" | "j = j2" | "j = j2 + 1" | "j = j2 + 2" | "j = j2 + 3" | "j = j2 + 4" | "j = j2 + 5" | "j \ j1 \ j \ j2 \ j \ j2 + 1 \ j \ j2 + 2 \ j \ j2 + 3 \ j \ j2 + 4 \ j \ j2 + 5" by auto then show ?thesis using tpsB2_def tpsB3_def tps0 jk by (cases, simp_all only: nth_list_update_eq nth_list_update_neq length_list_update, metis nth_list_update_neq) qed qed qed definition "tpsI0 t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 4 := (\1\\<^sub>N, 1), j2 + 5 := (\rv t\\<^sub>N, 1)]" definition "tpsI1 t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 5 := (\rv t\\<^sub>N, 1)]" lemma tmI1 [transforms_intros]: assumes "t < idx" and "ns ! idx = ns ! t" shows "transforms tmI1 (tpsB3 t) 12 (tpsI1 t)" unfolding tmI1_def proof (tform tps: tpsB3_def jk assms(2) time: assms nlength_1_simp) show "tpsI1 t = (tpsB3 t)[j2 + 4 := (\0\\<^sub>N, 1)]" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using tpsB3_def tpsI1_def tps0 jk by simp show "?l ! j = ?r ! j" if "j < length ?l" for j proof - consider "j = j1" | "j = j2" | "j = j2 + 1" | "j = j2 + 2" | "j = j2 + 3" | "j = j2 + 4" | "j = j2 + 5" | "j \ j1 \ j \ j2 \ j \ j2 + 1 \ j \ j2 + 2 \ j \ j2 + 3 \ j \ j2 + 4 \ j \ j2 + 5" by auto then show ?thesis using tpsB3_def tpsI1_def tps0 jk by (cases, simp_all only: nth_list_update_eq nth_list_update_neq length_list_update, metis nth_list_update_neq) qed qed qed definition "tpsI2 t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 5 := (\t\\<^sub>N, 1)]" lemma tmI2 [transforms_intros]: assumes "ttt = 26 + 3 * nlength t + 3 * nlength (rv t)" and "t < idx" and "ns ! idx = ns ! t" shows "transforms tmI2 (tpsB3 t) ttt (tpsI2 t)" unfolding tmI2_def proof (tform tps: assms(2,3) tpsI1_def jk time: assms(1)) show "tpsI2 t = (tpsI1 t)[j2 + 5 := (\t\\<^sub>N, 1)]" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using tpsI1_def tpsI2_def tps0 by simp show "?l ! j = ?r ! j" if "j < length ?l" for j proof - consider "j = j1" | "j = j2" | "j = j2 + 1" | "j = j2 + 2" | "j = j2 + 3" | "j = j2 + 4" | "j = j2 + 5" | "j \ j1 \ j \ j2 \ j \ j2 + 1 \ j \ j2 + 2 \ j \ j2 + 3 \ j \ j2 + 4 \ j \ j2 + 5" by auto then show ?thesis using tpsI1_def tpsI2_def tps0 jk assms(2,3) by (cases, simp_all only: One_nat_def nth_list_update_eq nth_list_update_neq length_list_update list_update_overwrite) qed qed qed definition "tpsI t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\t\\<^sub>N, 1), j2 + 5 := (\rv (Suc t)\\<^sub>N, 1)]" lemma tmI [transforms_intros]: assumes "ttt = 28 + 6 * nllength ns" and "t < idx" shows "transforms tmI (tpsB3 t) ttt (tpsI t)" unfolding tmI_def proof (tform tps: assms) let ?tT = "26 + 3 * nlength t + 3 * nlength (rv t)" have *: "(ns ! idx = ns ! t) = (read (tpsB3 t) ! (j2 + 4) \ \)" using condition_tpsB3 by simp then show "read (tpsB3 t) ! (j2 + 4) \ \ \ ns ! idx = ns ! t" by simp have "ns ! idx = ns ! t \ rv (Suc t) = t" using rv_change idx assms(2) by simp then show "read (tpsB3 t) ! (j2 + 4) \ \ \ tpsI t = tpsI2 t" using tpsI_def tpsI2_def * by simp have "ns ! idx \ ns ! t \ rv (Suc t) = rv t" using rv_keep idx assms(2) by simp then have "ns ! idx \ ns ! t \ tpsI t = tpsB3 t" using tpsI_def tpsB3_def tps0 jk by (smt (verit, ccfv_SIG) add_left_imp_eq list_update_id list_update_swap numeral_eq_iff one_eq_numeral_iff semiring_norm(83) semiring_norm(87)) then show "\ read (tpsB3 t) ! (j2 + 4) \ \ \ tpsI t = tpsB3 t" using * by simp show "26 + 3 * nlength t + 3 * nlength (rv t) + 2 \ ttt" proof - have "26 + 3 * nlength t + 3 * nlength (rv t) + 2 = 28 + 3 * nlength t + 3 * nlength (rv t)" by simp also have "... \ 28 + 3 * nlength idx + 3 * nlength (rv t)" using assms(2) nlength_mono by simp also have "... \ 28 + 3 * nlength idx + 3 * nlength idx" proof - have "rv t \ idx" using assms(2) rv_le by (metis less_or_eq_imp_le max_absorb2) then show ?thesis using nlength_mono by simp qed also have "... = 28 + 6 * nlength idx" by simp also have "... \ 28 + 6 * nllength ns" proof - have "nlength idx \ nlength (length ns)" using idx nlength_mono by simp then have "nlength idx \ length ns" using nlength_le_n by (meson le_trans) then have "nlength idx \ nllength ns" using length_le_nllength le_trans by blast then show ?thesis by simp qed finally show ?thesis using assms(1) by simp qed qed lemma tmB4: assumes "ttt = 71 + 24 * (nllength ns)\<^sup>2 + 2 * nlength (ns ! t) + 6 * nllength ns" and "t < idx" shows "transforms tmB4 (tpsC t) ttt (tpsI t)" unfolding tmB4_def by (tform tps: assms(2) jk time: assms(1)) lemma tmB4' [transforms_intros]: assumes "ttt = 71 + 32 * (nllength ns)\<^sup>2" and "t < idx" shows "transforms tmB4 (tpsC t) ttt (tpsI t)" proof - let ?ttt = "71 + 24 * (nllength ns)\<^sup>2 + 2 * nlength (ns ! t) + 6 * nllength ns" have "?ttt \ 71 + 24 * (nllength ns)\<^sup>2 + 2 * nlength (ns ! t) + 6 * nllength ns ^ 2" by (metis add_mono_thms_linordered_semiring(2) le_square mult.commute mult_le_mono1 power2_eq_square) also have "... = 71 + 30 * (nllength ns)\<^sup>2 + 2 * nlength (ns ! t)" by simp also have "... \ 71 + 30 * (nllength ns)\<^sup>2 + 2 * nllength ns" using member_le_nllength assms(2) idx by simp also have "... \ 71 + 30 * (nllength ns)\<^sup>2 + 2 * nllength ns ^ 2" by (simp add: power2_eq_square) also have "... = 71 + 32 * (nllength ns)\<^sup>2" by simp finally have "?ttt \ 71 + 32 * (nllength ns)\<^sup>2" . then show ?thesis using assms tmB4 transforms_monotone by blast qed definition "tpsB5 t \ tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\Suc t\\<^sub>N, 1), j2 + 5 := (\rv (Suc t)\\<^sub>N, 1)]" lemma tmB5: assumes "ttt = 76 + 32 * (nllength ns)\<^sup>2 + 2 * nlength t" and "t < idx" shows "transforms tmB5 (tpsC t) ttt (tpsB5 t)" unfolding tmB5_def proof (tform tps: assms(2) tpsI_def jk time: assms(1)) show "tpsB5 t = (tpsI t)[j2 + 2 := (\Suc t\\<^sub>N, 1)]" using tpsB5_def tpsI_def - by (smt (z3) add_left_imp_eq list_update_overwrite list_update_swap numeral_eq_iff semiring_norm(89)) + by (smt (verit) add_left_imp_eq list_update_overwrite list_update_swap numeral_eq_iff semiring_norm(89)) qed lemma tmB5' [transforms_intros]: assumes "ttt = 76 + 34 * (nllength ns)\<^sup>2" and "t < idx" shows "transforms tmB5 (tpsC t) ttt (tpsL (Suc t))" proof - have "76 + 32 * (nllength ns)\<^sup>2 + 2 * nlength t \ 76 + 32 * (nllength ns)\<^sup>2 + 2 * nllength ns" using assms(2) idx length_le_nllength nlength_le_n by (meson add_mono_thms_linordered_semiring(2) le_trans less_or_eq_imp_le mult_le_mono2) also have "... \ 76 + 32 * (nllength ns)\<^sup>2 + 2 * nllength ns ^ 2" by (simp add: power2_eq_square) also have "... \ 76 + 34 * (nllength ns)\<^sup>2" by simp finally have "76 + 32 * (nllength ns)\<^sup>2 + 2 * nlength t \ 76 + 34 * (nllength ns)\<^sup>2" . moreover have "tpsL (Suc t) = tpsB5 t" using tpsL_def tpsB5_def by simp ultimately show ?thesis using assms tmB5 transforms_monotone by fastforce qed lemma tmL [transforms_intros]: assumes "ttt = 125 * nllength ns ^ 3 + 8" and "iidx = idx" shows "transforms tmL (tpsL 0) ttt (tpsC iidx)" unfolding tmL_def proof (tform tps: assms) let ?tC = "3 * nllength ns ^ 2 + 7" let ?tB5 = "76 + 34 * (nllength ns)\<^sup>2" show "\t. t < iidx \ read (tpsC t) ! (j2 + 4) = \" using condition_tpsC assms(2) by fast show "read (tpsC iidx) ! (j2 + 4) \ \" using condition_tpsC assms(2) by simp have "iidx * (?tC + ?tB5 + 2) + ?tC + 1 = iidx * (37 * (nllength ns)\<^sup>2 + 85) + 3 * (nllength ns)\<^sup>2 + 8" by simp also have "... \ length ns * (37 * (nllength ns)\<^sup>2 + 85) + 3 * (nllength ns)\<^sup>2 + 8" using assms(2) idx by simp also have "... \ nllength ns * (37 * (nllength ns)\<^sup>2 + 85) + 3 * (nllength ns)\<^sup>2 + 8" using length_le_nllength by simp also have "... = 37 * nllength ns ^ 3 + 85 * nllength ns + 3 * (nllength ns)\<^sup>2 + 8" by algebra also have "... \ 37 * nllength ns ^ 3 + 85 * nllength ns + 3 * nllength ns ^ 3 + 8" proof - have "nllength ns ^ 2 \ nllength ns ^ 3" by (metis Suc_eq_plus1 add.commute eq_refl le_add2 neq0_conv numeral_3_eq_3 numerals(2) pow_mono power_eq_0_iff zero_less_Suc) then show ?thesis by simp qed also have "... \ 37 * nllength ns ^ 3 + 85 * nllength ns ^ 3 + 3 * nllength ns ^ 3 + 8" by (simp add: power3_eq_cube) also have "... = 125 * nllength ns ^ 3 + 8" by simp finally have "iidx * (?tC + ?tB5 + 2) + ?tC + 1 \ 125 * nllength ns ^ 3 + 8" . then show "iidx * (?tC + ?tB5 + 2) + ?tC + 1 \ ttt" using assms(1) by simp qed lemma tm2' [transforms_intros]: assumes "ttt = 14 + 3 * nlength idx + (21 * (nllength ns)\<^sup>2 + 26)" shows "transforms tm2 tps0 ttt (tpsL 0)" using tm2 assms tpsL_eq_tps2 by simp lemma tm3 [transforms_intros]: assumes "ttt = 40 + (3 * nlength idx + 21 * (nllength ns)\<^sup>2) + (125 * nllength ns ^ 3 + 8)" shows "transforms tm3 tps0 ttt (tpsC idx)" unfolding tm3_def by (tform tps: assms jk) lemma tpsC_idx: "tpsC idx = tps0 [j2 + 1 := (\ns ! idx\\<^sub>N, 1), j2 + 2 := (\idx\\<^sub>N, 1), j2 + 4 := (\1\\<^sub>N, 1), j2 + 5 := (\if \i ns ! i = ns ! idx else idx\\<^sub>N, 1)]" using tpsC_def rv_def by simp definition tps4 :: "tape list" where "tps4 \ tps0 [j2 + 1 := (\[]\, 1), j2 + 2 := (\idx\\<^sub>N, 1), j2 + 4 := (\1\\<^sub>N, 1), j2 + 5 := (\if \i ns ! i = ns ! idx else idx\\<^sub>N, 1)]" lemma tm4 [transforms_intros]: assumes "ttt = 55 + 3 * nlength idx + 21 * (nllength ns)\<^sup>2 + 125 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: assms tpsC_def tps4_def jk) show "proper_symbols (canrepr (ns ! idx))" using proper_symbols_canrepr by simp show "tps4 = (tpsC idx)[j2 + 1 := (\[]\, 1)]" using tpsC_idx tps4_def by (simp add: list_update_swap[of "Suc j2"]) qed definition tps5 :: "tape list" where "tps5 \ tps0 [j2 + 1 := (\[]\, 1), j2 + 2 := (\[]\, 1), j2 + 4 := (\1\\<^sub>N, 1), j2 + 5 := (\if \i ns ! i = ns ! idx else idx\\<^sub>N, 1)]" lemma tm5 [transforms_intros]: assumes "ttt = 62 + 5 * nlength idx + 21 * (nllength ns)\<^sup>2 + 125 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def proof (tform tps: tps4_def tps5_def jk time: assms tps4_def jk) show "proper_symbols (canrepr idx)" using proper_symbols_canrepr by simp qed definition tps6 :: "tape list" where "tps6 \ tps0 [j2 + 1 := (\[]\, 1), j2 + 2 := (\[]\, 1), j2 + 4 := (\[]\, 1), j2 + 5 := (\if \i ns ! i = ns ! idx else idx\\<^sub>N, 1)]" lemma tm6: assumes "ttt = 71 + 5 * nlength idx + 21 * (nllength ns)\<^sup>2 + 125 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def proof (tform tps: tps5_def tps6_def jk time: assms tps5_def jk nlength_1_simp) show "proper_symbols (canrepr (Suc 0))" using proper_symbols_canrepr by simp qed definition tps6' :: "tape list" where "tps6' \ tps0 [j2 + 5 := (\if \i ns ! i = ns ! idx else idx\\<^sub>N, 1)]" lemma tps6'_eq_tps6: "tps6' = tps6" using tps6'_def tps6_def tps0 jk canrepr_0 by (metis (no_types, lifting) list_update_id) lemma tm6': assumes "ttt = 71 + 153 * nllength ns ^ 3" shows "transforms tm6 tps0 ttt tps6'" proof - let ?ttt = "71 + 5 * nlength idx + 21 * (nllength ns)\<^sup>2 + 125 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" have "?ttt \ 71 + 5 * nlength idx + 21 * (nllength ns)^3 + 125 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" using pow_mono[of 2 3 "nllength ns"] by fastforce also have "... = 71 + 5 * nlength idx + 146 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" by simp also have "... \ 71 + 5 * nllength ns + 146 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" proof - have "nlength idx \ nllength ns" using idx by (meson le_trans length_le_nllength nlength_le_n order.strict_implies_order) then show ?thesis by simp qed also have "... \ 71 + 5 * nllength ns ^ 3 + 146 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" using linear_le_pow by simp also have "... = 71 + 151 * nllength ns ^ 3 + 2 * nlength (ns ! idx)" by simp also have "... \ 71 + 151 * nllength ns ^ 3 + 2 * nllength ns" using idx member_le_nllength by simp also have "... \ 71 + 151 * nllength ns ^ 3 + 2 * nllength ns^3" using linear_le_pow by simp also have "... = 71 + 153 * nllength ns ^ 3" by simp finally have "?ttt \ ttt" using assms by simp then have "transforms tm6 tps0 ttt tps6" using tm6 transforms_monotone by simp then show ?thesis using tps6'_eq_tps6 by simp qed end (* context tps0 idx ns *) end (* locale turing_machine_prev *) lemma transforms_tm_prevI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k idx :: nat and ns :: "nat list" assumes "0 < j1" "j1 < j2" "j2 + 6 \ k" "length tps = k" and "idx < length ns" assumes "tps ! j1 = (\ns\\<^sub>N\<^sub>L, 1)" "tps ! j2 = (\idx\\<^sub>N, 1)" "tps ! (j2 + 1) = (\0\\<^sub>N, 1)" "tps ! (j2 + 2) = (\0\\<^sub>N, 1)" "tps ! (j2 + 3) = (\0\\<^sub>N, 1)" "tps ! (j2 + 4) = (\0\\<^sub>N, 1)" "tps ! (j2 + 5) = (\0\\<^sub>N, 1)" assumes "ttt = 71 + 153 * nllength ns ^ 3" assumes "tps' = tps [j2 + 5 := (\previous ns idx\\<^sub>N, 1)]" shows "transforms (tm_prev j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_prev j1 j2 . show ?thesis using assms loc.tm6_eq_tm_prev loc.tm6' loc.tps6'_def previous_def by simp qed subsection \Checking containment in a list\ text \ A Turing machine that checks whether a number given on tape $j_2$ is contained in the list of numbers on tape $j_1$. If so, it writes a~1 to tape $j_3$, and otherwise leaves tape $j_3$ unchanged. So tape $j_3$ should be initialized with~0. \ definition tm_contains :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_contains j1 j2 j3 \ WHILE [] ; \rs. rs ! j1 \ \ DO tm_nextract \ j1 (j3 + 1) ;; tm_equalsn j2 (j3 + 1) (j3 + 2) ;; IF \rs. rs ! (j3 + 2) \ \ THEN tm_setn j3 1 ELSE [] ENDIF ;; tm_setn (j3 + 1) 0 ;; tm_setn (j3 + 2) 0 DONE ;; tm_cr j1" lemma tm_contains_tm: assumes "0 < j1" "j1 \ j2" "j3 + 2 < k" "j1 < j3" "j2 < j3" and "k \ 2" and "G \ 5" shows "turing_machine k G (tm_contains j1 j2 j3)" unfolding tm_contains_def using tm_nextract_tm tm_equalsn_tm Nil_tm tm_setn_tm tm_cr_tm assms turing_machine_branch_turing_machine turing_machine_loop_turing_machine by simp locale turing_machine_contains = fixes j1 j2 j3 :: tapeidx begin definition "tmL1 \ tm_nextract \ j1 (j3 + 1)" definition "tmL2 \ tmL1 ;; tm_equalsn j2 (j3 + 1) (j3 + 2)" definition "tmI \ IF \rs. rs ! (j3 + 2) \ \ THEN tm_setn j3 1 ELSE [] ENDIF" definition "tmL3 \ tmL2 ;; tmI" definition "tmL4 \ tmL3 ;; tm_setn (j3 + 1) 0" definition "tmL5 \ tmL4 ;; tm_setn (j3 + 2) 0" definition "tmL \ WHILE [] ; \rs. rs ! j1 \ \ DO tmL5 DONE" definition "tm2 \ tmL ;; tm_cr j1" lemma tm2_eq_tm_contains: "tm2 = tm_contains j1 j2 j3" unfolding tm2_def tmL_def tmL5_def tmL4_def tmL3_def tmI_def tmL2_def tmL1_def tm_contains_def by simp context fixes tps0 :: "tape list" and k :: nat and ns :: "nat list" and needle :: nat assumes jk: "0 < j1" "j1 \ j2" "j3 + 2 < k" "j1 < j3" "j2 < j3" "length tps0 = k" and tps0: "tps0 ! j1 = nltape' ns 0" "tps0 ! j2 = (\needle\\<^sub>N, 1)" "tps0 ! j3 = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 2) = (\0\\<^sub>N, 1)" begin definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j1 := nltape' ns t, j3 := (\\i\<^sub>B, 1)]" lemma tpsL0: "tpsL 0 = tps0" unfolding tpsL_def using tps0 by (smt (verit) list_update_id not_less_zero) definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0 [j1 := nltape' ns (Suc t), j3 := (\\i\<^sub>B, 1), j3 + 1 := (\ns ! t\\<^sub>N, 1)]" lemma tmL1 [transforms_intros]: assumes "ttt = 12 + 2 * nlength (ns ! t)" and "t < length ns" shows "transforms tmL1 (tpsL t) ttt (tpsL1 t)" unfolding tmL1_def proof (tform tps: assms(2) tpsL_def tpsL1_def jk tps0) show "ttt = 12 + 2 * nlength 0 + 2 * nlength (ns ! t)" using assms(1) by simp show "tpsL1 t = (tpsL t) [j1 := nltape' ns (Suc t), j3 + 1 := (\ns ! t\\<^sub>N, 1)]" unfolding tpsL1_def tpsL_def using jk by (simp add: list_update_swap) qed definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [j1 := nltape' ns (Suc t), j3 := (\\i\<^sub>B, 1), j3 + 1 := (\ns ! t\\<^sub>N, 1), j3 + 2 := (\needle = ns ! t\\<^sub>B, 1)]" lemma tmL2 [transforms_intros]: assumes "ttt = 12 + 2 * nlength (ns ! t) + (3 * nlength (min needle (ns ! t)) + 7)" and "t < length ns" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" unfolding tmL2_def by (tform tps: assms tps0 tpsL1_def tpsL2_def jk) definition tpsI :: "nat \ tape list" where "tpsI t \ tps0 [j1 := nltape' ns (Suc t), j3 := (\\i\<^sub>B, 1), j3 + 1 := (\ns ! t\\<^sub>N, 1), j3 + 2 := (\needle = ns ! t\\<^sub>B, 1)]" lemma tmI [transforms_intros]: assumes "ttt = 16" and "t < length ns" shows "transforms tmI (tpsL2 t) ttt (tpsI t)" unfolding tmI_def proof (tform tps: tpsL2_def jk) show "10 + 2 * nlength (if \i ttt" using assms(1) nlength_1_simp nlength_0 by simp show "0 + 1 \ ttt" using assms(1) by simp have "tpsL2 t ! (j3 + 2) = (\needle = ns ! t\\<^sub>B, 1)" using tpsL2_def jk by simp then have *: "(read (tpsL2 t) ! (j3 + 2) = \) = (needle \ ns ! t)" using jk read_ncontents_eq_0[of "tpsL2 t" "j3 + 2"] tpsL2_def by simp show "tpsI t = (tpsL2 t)[j3 := (\1\\<^sub>N, 1)]" if "read (tpsL2 t) ! (j3 + 2) \ \" proof - have "needle = ns ! t" using * that by simp then have "\i read (tpsL2 t) ! (j3 + 2) \ \" proof - have "needle \ ns ! t" using * that by simp then have "(\ii< t. ns ! i = needle)" using assms(2) less_Suc_eq by auto then show ?thesis unfolding tpsI_def tpsL2_def by simp qed qed lemma tmL3 [transforms_intros]: assumes "ttt = 28 + 2 * nlength (ns ! t) + (3 * nlength (min needle (ns ! t)) + 7)" and "t < length ns" shows "transforms tmL3 (tpsL t) ttt (tpsI t)" unfolding tmL3_def by (tform tps: assms) definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [j1 := nltape' ns (Suc t), j3 := (\\i\<^sub>B, 1), j3 + 1 := (\0\\<^sub>N, 1), j3 + 2 := (\needle = ns ! t\\<^sub>B, 1)]" lemma tmL4 [transforms_intros]: assumes "ttt = 38 + 4 * nlength (ns ! t) + (3 * nlength (min needle (ns ! t)) + 7)" and "t < length ns" shows "transforms tmL4 (tpsL t) ttt (tpsL4 t)" unfolding tmL4_def proof (tform tps: assms tpsI_def jk) show "tpsL4 t = (tpsI t)[j3 + 1 := (\0\\<^sub>N, 1)]" unfolding tpsL4_def tpsI_def by (simp add: list_update_swap) qed definition tpsL5 :: "nat \ tape list" where "tpsL5 t \ tps0 [j1 := nltape' ns (Suc t), j3 := (\\i\<^sub>B, 1), j3 + 1 := (\0\\<^sub>N, 1), j3 + 2 := (\0\\<^sub>N, 1)]" lemma tmL5: assumes "ttt = 48 + 4 * nlength (ns ! t) + (3 * nlength (min needle (ns ! t)) + 7) + 2 * nlength (if needle = ns ! t then 1 else 0)" and "t < length ns" shows "transforms tmL5 (tpsL t) ttt (tpsL5 t)" unfolding tmL5_def proof (tform tps: assms tpsL4_def jk) show "tpsL5 t = (tpsL4 t)[j3 + 2 := (\0\\<^sub>N, 1)]" unfolding tpsL5_def tpsL4_def by (simp add: list_update_swap) qed definition tpsL5' :: "nat \ tape list" where "tpsL5' t \ tps0 [j1 := nltape' ns (Suc t), j3 := (\\i\<^sub>B, 1)]" lemma tpsL5': "tpsL5' t = tpsL5 t" using tpsL5'_def tpsL5_def tps0 jk by (smt (verit, del_insts) One_nat_def less_Suc_eq less_add_same_cancel1 list_update_swap not_less_eq tape_list_eq zero_less_numeral) lemma tmL5': assumes "ttt = 57 + 4 * nlength (ns ! t) + 3 * nlength (min needle (ns ! t))" and "t < length ns" shows "transforms tmL5 (tpsL t) ttt (tpsL5' t)" proof - have "nlength (if needle = ns ! t then 1 else 0) \ 1" using nlength_1_simp by simp then have "48 + 4 * nlength (ns ! t) + (3 * nlength (min needle (ns ! t)) + 7) + 2 * nlength (if needle = ns ! t then 1 else 0) \ ttt" using assms(1) by simp then show ?thesis using tpsL5' tmL5 transforms_monotone assms(2) by fastforce qed lemma tmL5'' [transforms_intros]: assumes "ttt = 57 + 7 * nllength ns" and "t < length ns" shows "transforms tmL5 (tpsL t) ttt (tpsL (Suc t))" proof - have "nlength (ns ! t) \ nllength ns" using assms(2) by (simp add: member_le_nllength) moreover from this have "nlength (min needle (ns ! t)) \ nllength ns" using nlength_mono by (metis dual_order.trans min_def) ultimately have "ttt \ 57 + 4 * nlength (ns ! t) + 3 * nlength (min needle (ns ! t))" using assms(1) by simp moreover have "tpsL5' t = tpsL (Suc t)" using tpsL5'_def tpsL_def by simp ultimately show ?thesis using tmL5' assms(2) transforms_monotone by fastforce qed lemma tmL [transforms_intros]: assumes "ttt = length ns * (59 + 7 * nllength ns) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length ns))" unfolding tmL_def proof (tform) let ?t = "57 + 7 * nllength ns" show "length ns * (57 + 7 * nllength ns + 2) + 1 \ ttt" using assms by simp have *: "tpsL t ! j1 = nltape' ns t" for t using tpsL_def jk by simp moreover have "read (tpsL t) ! j1 = tpsL t :.: j1" for t using tapes_at_read'[of j1 "tpsL t"] tpsL_def jk by simp ultimately have "read (tpsL t) ! j1 = |.| (nltape' ns t)" for t by simp then have "read (tpsL t) ! j1 = \ \ (t \ length ns)" for t using nltape'_tape_read by simp then show "\i. i < length ns \ read (tpsL i) ! j1 \ \" "\ read (tpsL (length ns)) ! j1 \ \" using * by simp_all qed definition tps2 :: "tape list" where "tps2 \ tps0 [j1 := nltape' ns 0, j3 := (\\i\<^sub>B, 1)]" lemma tm2: assumes "ttt = length ns * (59 + 7 * nllength ns) + nllength ns + 4" shows "transforms tm2 (tpsL 0) ttt tps2" unfolding tm2_def proof (tform tps: tpsL_def jk) show "clean_tape (tpsL (length ns) ! j1)" using tpsL_def jk clean_tape_nlcontents by simp have "tpsL (length ns) ! j1 |#=| 1 = nltape' ns 0" using tpsL_def jk by simp then have "(tpsL (length ns))[j1 := tpsL (length ns) ! j1 |#=| 1] = (tpsL (length ns))[j1 := nltape' ns 0]" by simp then show "tps2 = (tpsL (length ns))[j1 := tpsL (length ns) ! j1 |#=| 1]" unfolding tps2_def tpsL_def using jk by (simp add: list_update_swap) have "tpsL (length ns) :#: j1 = Suc (nllength ns)" using tpsL_def jk by simp then show "ttt = length ns * (59 + 7 * nllength ns) + 1 + (tpsL (length ns) :#: j1 + 2)" using assms by simp qed definition tps2' :: "tape list" where "tps2' \ tps0 [j3 := (\needle \ set ns\\<^sub>B, 1)]" lemma tm2': assumes "ttt = 67 * nllength ns ^ 2 + 4" shows "transforms tm2 tps0 ttt tps2'" proof - let ?t = "length ns * (59 + 7 * nllength ns) + nllength ns + 4" have "?t \ nllength ns * (59 + 7 * nllength ns) + nllength ns + 4" by (simp add: length_le_nllength) also have "... = 60 * nllength ns + 7 * nllength ns ^ 2 + 4" by algebra also have "... \ 60 * nllength ns ^ 2 + 7 * nllength ns ^ 2 + 4" using linear_le_pow by simp also have "... = 67 * nllength ns ^ 2 + 4" by simp finally have "?t \ 67 * nllength ns ^ 2 + 4" . moreover have "tps2' = tps2" unfolding tps2_def tps2'_def using tps0(1) by (smt (verit) in_set_conv_nth list_update_id) ultimately show ?thesis using tm2 assms transforms_monotone tpsL0 by simp qed end (* context *) end (* locale *) lemma transforms_tm_containsI [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and ttt k needle :: nat and ns :: "nat list" assumes "0 < j1" "j1 \ j2" "j3 + 2 < k" "j1 < j3" "j2 < j3" "length tps = k" assumes "tps ! j1 = nltape' ns 0" "tps ! j2 = (\needle\\<^sub>N, 1)" "tps ! j3 = (\0\\<^sub>N, 1)" "tps ! (j3 + 1) = (\0\\<^sub>N, 1)" "tps ! (j3 + 2) = (\0\\<^sub>N, 1)" assumes "ttt = 67 * nllength ns ^ 2 + 4" assumes "tps' = tps [j3 := (\needle \ set ns\\<^sub>B, 1)]" shows "transforms (tm_contains j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_contains j1 j2 j3 . show ?thesis using assms loc.tm2_eq_tm_contains loc.tps2'_def loc.tm2' by simp qed subsection \Creating lists of consecutive numbers\ text \ The next TM accepts a number $start$ on tape $j_1$ and a number $delta$ on tape $j_2$. It outputs the list $[start, \dots, start + delta - 1]$ on tape $j_3$. \ definition tm_range :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_range j1 j2 j3 \ tm_copyn j1 (j3 + 2) ;; tm_copyn j2 (j3 + 1) ;; WHILE [] ; \rs. rs ! (j3 + 1) \ \ DO tm_append j3 (j3 + 2) ;; tm_incr (j3 + 2) ;; tm_decr (j3 + 1) DONE ;; tm_setn (j3 + 2) 0 ;; tm_cr j3" lemma tm_range_tm: assumes "k \ j3 + 3" and "G \ 5" and "j1 \ j2" and "0 < j1" and "0 < j2" and "j1 < j3" and "j2 < j3" shows "turing_machine k G (tm_range j1 j2 j3)" unfolding tm_range_def using assms tm_copyn_tm tm_decr_tm tm_append_tm tm_setn_tm tm_incr_tm Nil_tm tm_cr_tm turing_machine_loop_turing_machine by simp locale turing_machine_range = fixes j1 j2 j3 :: tapeidx begin definition "tm1 \ tm_copyn j1 (j3 + 2)" definition "tm2 \ tm1 ;; tm_copyn j2 (j3 + 1)" definition "tmB1 \ tm_append j3 (j3 + 2)" definition "tmB2 \ tmB1 ;; tm_incr (j3 + 2)" definition "tmB3 \ tmB2 ;; tm_decr (j3 + 1)" definition "tmL \ WHILE [] ; \rs. rs ! (j3 + 1) \ \ DO tmB3 DONE" definition "tm3 \ tm2 ;; tmL" definition "tm4 \ tm3 ;; tm_setn (j3 + 2) 0" definition "tm5 \ tm4 ;; tm_cr j3" lemma tm5_eq_tm_range: "tm5 = tm_range j1 j2 j3" unfolding tm_range_def tm5_def tm4_def tm3_def tmL_def tmB3_def tmB2_def tmB1_def tm2_def tm1_def by simp context fixes tps0 :: "tape list" and k start delta :: nat assumes jk: "k \ j3 + 3" "j1 \ j2" "0 < j1" "0 < j2" "0 < j3" "j1 < j3" "j2 < j3" "length tps0 = k" and tps0: "tps0 ! j1 = (\start\\<^sub>N, 1)" "tps0 ! j2 = (\delta\\<^sub>N, 1)" "tps0 ! j3 = (\[]\\<^sub>N\<^sub>L, 1)" "tps0 ! (j3 + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 2) = (\0\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j3 + 2 := (\start\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 14 + 3 * nlength start" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def by (tform tps: nlength_0 assms tps0 tps1_def jk) definition "tps2 \ tps0 [j3 + 2 := (\start\\<^sub>N, 1), j3 + 1 := (\delta\\<^sub>N, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 28 + 3 * nlength start + 3 * nlength delta" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: jk tps0 tps1_def tps2_def) show "ttt = 14 + 3 * nlength start + (14 + 3 * (nlength delta + nlength 0))" using assms by simp qed definition "tpsL t \ tps0 [j3 + 2 := (\start + t\\<^sub>N, 1), j3 + 1 := (\delta - t\\<^sub>N, 1), j3 := nltape [start.. tps0 [j3 + 2 := (\start + t\\<^sub>N, 1), j3 + 1 := (\delta - t\\<^sub>N, 1), j3 := nltape ([start.. tps0 [j3 + 2 := (\Suc (start + t)\\<^sub>N, 1), j3 + 1 := (\delta - t\\<^sub>N, 1), j3 := nltape ([start.. tps0 [j3 + 2 := (\Suc (start + t)\\<^sub>N, 1), j3 + 1 := (\delta - t - 1\\<^sub>N, 1), j3 := nltape ([start.. tpsL (Suc t)" using tpsB3_def tpsL_def by simp lemma tmB3' [transforms_intros]: assumes "ttt = 19 + 6 * nlength (start + delta)" and "t < delta" shows "transforms tmB3 (tpsL t) ttt (tpsL (Suc t))" proof - have "19 + 4 * nlength (start + t) + 2 * nlength (delta - t) \ 19 + 4 * nlength (start + t) + 2 * nlength delta" using nlength_mono by simp also have "... \ 19 + 4 * nlength (start + delta) + 2 * nlength delta" using assms(2) nlength_mono by simp also have "... \ 19 + 4 * nlength (start + delta) + 2 * nlength (start + delta)" using nlength_mono by simp also have "... = 19 + 6 * nlength (start + delta)" by simp finally have "19 + 4 * nlength (start + t) + 2 * nlength (delta - t) \ 19 + 6 * nlength (start + delta)" . then show ?thesis using tpsB3 tmB3 transforms_monotone assms(1) by metis qed lemma tmL: assumes "ttt = delta * (21 + 6 * nlength (start + delta)) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL delta)" unfolding tmL_def proof (tform) have "read (tpsL t) ! (j3 + 1) \ \ \ t < delta" for t using tpsL_def read_ncontents_eq_0 jk by auto then show "\t. t < delta \ read (tpsL t) ! (j3 + 1) \ \" and "\ read (tpsL delta) ! (j3 + 1) \ \" by simp_all show "delta * (19 + 6 * nlength (start + delta) + 2) + 1 \ ttt" using assms(1) by simp qed lemma tmL' [transforms_intros]: assumes "ttt = delta * (21 + 6 * nlength (start + delta)) + 1" shows "transforms tmL tps2 ttt (tpsL delta)" using tmL assms tpsL_eq_tps2 by simp definition "tps3 \ tps0 [j3 + 2 := (\start + delta\\<^sub>N, 1), j3 := nltape [start.. tps0 [j3 := nltape [start..0\\<^sub>N, 1)]" using tps4_def tps3_def tps0 jk by (metis (mono_tags, lifting) Suc_neq_Zero add_cancel_right_right list_update_id list_update_overwrite list_update_swap numeral_2_eq_2) qed lemma tm4' [transforms_intros]: assumes "ttt = Suc delta * (39 + 8 * nlength (start + delta))" shows "transforms tm4 tps0 ttt tps4" proof - let ?ttt = "39 + 3 * nlength start + 3 * nlength delta + delta * (21 + 6 * nlength (start + delta)) + 2 * nlength (start + delta)" have "?ttt \ 39 + 3 * nlength (start + delta) + 3 * nlength (start + delta) + delta * (21 + 6 * nlength (start + delta)) + 2 * nlength (start + delta)" using nlength_mono by (meson add_le_mono add_le_mono1 le_add2 nat_add_left_cancel_le nat_le_iff_add nat_mult_le_cancel_disj) also have "... = 39 + 8 * nlength (start + delta) + delta * (21 + 6 * nlength (start + delta))" by simp also have "... \ 39 + 8 * nlength (start + delta) + delta * (39 + 8 * nlength (start + delta))" by simp also have "... = Suc delta * (39 + 8 * nlength (start + delta))" by simp finally have "?ttt \ Suc delta * (39 + 8 * nlength (start + delta))" . then show ?thesis using assms tm4 transforms_monotone tps4_def by simp qed definition "tps5 \ tps0 [j3 := (\[start..\<^sub>N\<^sub>L, 1)]" lemma tm5: assumes "ttt = Suc delta * (39 + 8 * nlength (start + delta)) + Suc (Suc (Suc (nllength [start.. Suc (nlength (start + delta)) * delta" using nllength_le_len_mult_max[of "[start.. Suc delta * (39 + 8 * nlength (start + delta)) + 3 + Suc (nlength (start + delta)) * delta" by simp also have "... \ 3 + Suc delta * (39 + 8 * nlength (start + delta)) + Suc delta * Suc (nlength (start + delta))" by simp also have "... = 3 + Suc delta * (39 + 8 * nlength (start + delta) + Suc (nlength (start + delta)))" by algebra also have "... = 3 + Suc delta * (40 + 9 * nlength (start + delta))" by simp also have "... \ Suc delta * (43 + 9 * nlength (start + delta))" by simp finally have "?ttt \ Suc delta * (43 + 9 * nlength (start + delta))" . then show ?thesis using tm5 assms transforms_monotone by simp qed end (* context tps0 *) end (* locale turing_machine_range *) lemma transforms_tm_rangeI [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and k start delta :: nat assumes "k \ j3 + 3" "j1 \ j2" "0 < j1" "0 < j2" "j1 < j3" "j2 < j3" "length tps = k" assumes "tps ! j1 = (\start\\<^sub>N, 1)" "tps ! j2 = (\delta\\<^sub>N, 1)" "tps ! j3 = (\[]\\<^sub>N\<^sub>L, 1)" "tps ! (j3 + 1) = (\0\\<^sub>N, 1)" "tps ! (j3 + 2) = (\0\\<^sub>N, 1)" assumes "ttt = Suc delta * (43 + 9 * nlength (start + delta))" assumes "tps' = tps [j3 := (\[start..\<^sub>N\<^sub>L, 1)]" shows "transforms (tm_range j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_range j1 j2 j3 . show ?thesis using assms loc.tm5_eq_tm_range loc.tm5' loc.tps5_def by simp qed subsection \Creating singleton lists\ text \ The next Turing machine appends the symbol \textbf{|} to the symbols on tape~$j$. Thus it turns a number into a singleton list containing this number. \ definition tm_to_list :: "tapeidx \ machine" where "tm_to_list j \ tm_right_until j {\} ;; tm_write j \ ;; tm_cr j" lemma tm_to_list_tm: assumes "0 < j" and "j < k" and "G \ 5" shows "turing_machine k G (tm_to_list j)" unfolding tm_to_list_def using tm_right_until_tm tm_write_tm tm_cr_tm assms by simp locale turing_machine_to_list = fixes j :: tapeidx begin definition "tm1 \ tm_right_until j {\}" definition "tm2 \ tm1 ;; tm_write j \" definition "tm3 \ tm2 ;; tm_cr j" lemma tm3_eq_tm_to_list: "tm3 = tm_to_list j" using tm3_def tm2_def tm1_def tm_to_list_def by simp context fixes tps0 :: "tape list" and k n :: nat assumes jk: "0 < j" "j < k" "length tps0 = k" and tps0: "tps0 ! j = (\n\\<^sub>N, 1)" begin definition "tps1 \ tps0[j := (\n\\<^sub>N, Suc (nlength n))]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (nlength n)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: assms tps1_def tps0 jk) show "rneigh (tps0 ! j) {0} (nlength n)" proof (rule rneighI) show "(tps0 ::: j) (tps0 :#: j + nlength n) \ {0}" using tps0 jk by simp show "\n'. n' < nlength n \ (tps0 ::: j) (tps0 :#: j + n') \ {0}" using assms tps0 jk bit_symbols_canrepr contents_def by fastforce qed qed definition "tps2 \ tps0[j := (\[n]\\<^sub>N\<^sub>L, Suc (nlength n))]" lemma tm2 [transforms_intros]: assumes "ttt = Suc (Suc (nlength n))" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: assms tps1_def tps0 jk) have "numlist [n] = canrepr n @ [\]" using numlist_def by simp then show "tps2 = tps1[j := tps1 ! j |:=| \]" using assms tps1_def tps2_def tps0 jk numlist_def nlcontents_def contents_snoc by simp qed definition "tps3 \ tps0[j := (\[n]\\<^sub>N\<^sub>L, 1)]" lemma tm3: assumes "ttt = 5 + 2 * nlength n" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps2_def tps0 jk time: assms tps2_def jk) show "clean_tape (tps2 ! j)" using tps2_def jk clean_tape_nlcontents by simp show "tps3 = tps2[j := tps2 ! j |#=| 1]" using tps3_def tps2_def jk by simp qed end (* context tps0 *) end (* locale turing_machine_tm_to_list *) lemma transforms_tm_to_listI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k n :: nat assumes "0 < j" "j < k" "length tps = k" assumes "tps ! j = (\n\\<^sub>N, 1)" assumes "ttt = 5 + 2 * nlength n" assumes "tps' = tps[j := (\[n]\\<^sub>N\<^sub>L, 1)]" shows "transforms (tm_to_list j) tps ttt tps'" proof - interpret loc: turing_machine_to_list j . show ?thesis using assms loc.tm3_eq_tm_to_list loc.tm3 loc.tps3_def by simp qed subsection \Extending with a list\ text \ The next Turing machine extends the list on tape $j_1$ with the list on tape $j_2$. We assume that the tape head on $j_1$ is already at the end of the list. \ definition tm_extend :: "tapeidx \ tapeidx \ machine" where "tm_extend j1 j2 \ tm_cp_until j2 j1 {\} ;; tm_cr j2" lemma tm_extend_tm: assumes "0 < j1" and "G \ 4" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_extend j1 j2)" unfolding tm_extend_def using assms tm_cp_until_tm tm_cr_tm by simp locale turing_machine_extend = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_cp_until j2 j1 {\}" definition "tm2 \ tm1 ;; tm_cr j2" lemma tm2_eq_tm_extend: "tm2 = tm_extend j1 j2" unfolding tm2_def tm2_def tm1_def tm_extend_def by simp context fixes tps0 :: "tape list" and k :: nat and ns1 ns2 :: "nat list" assumes jk: "0 < j1" "j1 < k" "j2 < k" "j1 \ j2" "length tps0 = k" assumes tps0: "tps0 ! j1 = nltape ns1" "tps0 ! j2 = (\ns2\\<^sub>N\<^sub>L, 1)" begin definition "tps1 \ tps0 [j1 := nltape (ns1 @ ns2), j2 := nltape ns2]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (nllength ns2)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps1_def tps0 jk) let ?n = "nllength ns2" show "rneigh (tps0 ! j2) {0} ?n" proof (rule rneighI) show "(tps0 ::: j2) (tps0 :#: j2 + nllength ns2) \ {0}" using tps0 nlcontents_def nllength_def jk by simp show "\i. i < nllength ns2 \ (tps0 ::: j2) (tps0 :#: j2 + i) \ {0}" using tps0 jk contents_def nlcontents_def nllength_def proper_symbols_numlist by fastforce qed show "ttt = Suc (nllength ns2)" using assms . show "tps1 = tps0 [j2 := tps0 ! j2 |+| nllength ns2, j1 := implant (tps0 ! j2) (tps0 ! j1) (nllength ns2)]" proof - have "implant (\ns2\\<^sub>N\<^sub>L, 1) (nltape ns1) (nllength ns2) = nltape (ns1 @ ns2)" using nlcontents_def nllength_def implant_contents by (simp add: numlist_append) moreover have "tps1 ! j2 = tps0 ! j2 |+| nllength ns2" using tps0 jk tps1_def by simp ultimately show ?thesis using tps0 jk tps1_def by (metis fst_conv list_update_swap plus_1_eq_Suc snd_conv) qed qed definition "tps2 \ tps0[j1 := nltape (ns1 @ ns2)]" lemma tm2: assumes "ttt = 4 + 2 * nllength ns2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps0 tps2_def tps1_def jk) show "clean_tape (tps1 ! j2)" using tps1_def jk clean_tape_nlcontents by simp show "ttt = Suc (nllength ns2) + (tps1 :#: j2 + 2)" using assms tps1_def jk by simp show "tps2 = tps1[j2 := tps1 ! j2 |#=| 1]" using tps1_def jk tps2_def tps0(2) by (metis fst_conv length_list_update list_update_id list_update_overwrite nth_list_update) qed end (* context tps0 *) end (* locale turing_machine_extend *) lemma transforms_tm_extendI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k :: nat and ns1 ns2 :: "nat list" assumes "0 < j1" "j1 < k" "j2 < k" "j1 \ j2" "length tps = k" assumes "tps ! j1 = nltape ns1" "tps ! j2 = (\ns2\\<^sub>N\<^sub>L, 1)" assumes "ttt = 4 + 2 * nllength ns2" assumes "tps' = tps[j1 := nltape (ns1 @ ns2)]" shows "transforms (tm_extend j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_extend j1 j2 . show ?thesis using loc.tm2_eq_tm_extend loc.tm2 loc.tps2_def assms by simp qed text \ An enhanced version of the previous Turing machine, the next one erases the list on tape $j_2$ after appending it to tape $j_1$. \ definition tm_extend_erase :: "tapeidx \ tapeidx \ machine" where "tm_extend_erase j1 j2 \ tm_extend j1 j2 ;; tm_erase_cr j2" lemma tm_extend_erase_tm: assumes "0 < j1" and "0 < j2" and "G \ 4" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_extend_erase j1 j2)" unfolding tm_extend_erase_def using assms tm_extend_tm tm_erase_cr_tm by simp lemma transforms_tm_extend_eraseI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k :: nat and ns1 ns2 :: "nat list" assumes "0 < j1" "j1 < k" "j2 < k" "j1 \ j2" "0 < j2" "length tps = k" assumes "tps ! j1 = nltape ns1" "tps ! j2 = (\ns2\\<^sub>N\<^sub>L, 1)" assumes "ttt = 11 + 4 * nllength ns2 " assumes "tps' = tps [j1 := nltape (ns1 @ ns2), j2 := (\[]\, 1)]" shows "transforms (tm_extend_erase j1 j2) tps ttt tps'" unfolding tm_extend_erase_def proof (tform tps: assms) let ?zs = "numlist ns2" show "tps[j1 := nltape (ns1 @ ns2)] ::: j2 = \?zs\" using assms by (simp add: nlcontents_def) show "proper_symbols ?zs" using proper_symbols_numlist by simp show "ttt = 4 + 2 * nllength ns2 + (tps[j1 := nltape (ns1 @ ns2)] :#: j2 + 2 * length (numlist ns2) + 6)" using assms nllength_def by simp qed section \Lists of lists of numbers\label{s:tm-numlistlist}\ text \ In this section we introduce a representation for lists of lists of numbers as symbol sequences over the quaternary alphabet @{text "\\\\"} and devise Turing machines for the few operations on such lists that we need. A tape containing such representations corresponds to a variable of type @{typ "nat list list"}. A tape in the start configuration corresponds to the empty list of lists of numbers. Many proofs in this section are copied from the previous section with minor modifications, mostly replacing the symbol @{text "\"} with @{text \}. \ subsection \Representation as symbol sequence\label{s:tm-numlistlist-repr}\ text \ We apply the same principle as for representing lists of numbers. To represent a list of lists of numbers, every element, which is now a list of numbers, is terminated by the symbol @{text \}. In this way the empty symbol sequence represents the empty list of lists of numbers. The list $[[]]$ containing only an empty list is represented by @{text \} and the list $[[0]]$ containing only a list containing only a~0 is represented by @{text "\\"}. As another example, the list $[[14,0,0,7],[],[0],[25]]$ is represented by @{text "\\\\\\\\\\\\\\\\\\\\\\"}. The number of @{text \} symbols equals the number of elements in the list. \null \ definition numlistlist :: "nat list list \ symbol list" where "numlistlist nss \ concat (map (\ns. numlist ns @ [\]) nss)" lemma numlistlist_Nil: "numlistlist [] = []" using numlistlist_def by simp proposition "numlistlist [[]] = [\]" using numlistlist_def by (simp add: numlist_Nil) lemma proper_symbols_numlistlist: "proper_symbols (numlistlist nss)" proof (induction nss) case Nil then show ?case using numlistlist_def by simp next case (Cons ns nss) have "numlistlist (ns # nss) = numlist ns @ [\] @ concat (map (\ns. numlist ns @ [\]) nss)" using numlistlist_def by simp then have "numlistlist (ns # nss) = numlist ns @ [\] @ numlistlist nss" using numlistlist_def by simp moreover have "proper_symbols (numlist ns)" using proper_symbols_numlist by simp moreover have "proper_symbols [\]" by simp ultimately show ?case using proper_symbols_append Cons by presburger qed lemma symbols_lt_append: fixes xs ys :: "symbol list" and z :: symbol assumes "symbols_lt z xs" and "symbols_lt z ys" shows "symbols_lt z (xs @ ys)" using assms prop_list_append by (simp add: nth_append) lemma symbols_lt_numlistlist: assumes "H \ 6" shows "symbols_lt H (numlistlist nss)" proof (induction nss) case Nil then show ?case using numlistlist_def by simp next case (Cons ns nss) have "numlistlist (ns # nss) = numlist ns @ [\] @ concat (map (\ns. numlist ns @ [\]) nss)" using numlistlist_def by simp then have "numlistlist (ns # nss) = numlist ns @ [\] @ numlistlist nss" using numlistlist_def by simp moreover have "symbols_lt H (numlist ns)" using assms numlist_234 nth_mem by fastforce moreover have "symbols_lt H [\]" using assms by simp ultimately show ?case using symbols_lt_append[of _ H] Cons by presburger qed lemma symbols_lt_prefix_eq: assumes "(x @ [\]) @ xs = (y @ [\]) @ ys" and "symbols_lt 5 x" and "symbols_lt 5 y" shows "x = y" proof - have *: "x @ [\] @ xs = y @ [\] @ ys" (is "?lhs = ?rhs") using assms(1) by simp show "x = y" proof (cases "length x \ length y") case True then have "?lhs ! i = ?rhs ! i" if "i < length x" for i using that * by simp then have eq: "x ! i = y ! i" if "i < length x" for i using that True by (metis Suc_le_eq le_trans nth_append) have "?lhs ! (length x) = \" by (metis Cons_eq_appendI nth_append_length) then have "?rhs ! (length x) = \" using * by metis moreover have "y ! i \ \" if "i < length y" for i using that assms(3) by auto ultimately have "length y \ length x" by (metis linorder_le_less_linear nth_append) then have "length y = length x" using True by simp then show ?thesis using eq by (simp add: list_eq_iff_nth_eq) next case False then have "?lhs ! i = ?rhs ! i" if "i < length y" for i using that * by simp have "?rhs ! (length y) = \" by (metis Cons_eq_appendI nth_append_length) then have "?lhs ! (length y) = \" using * by metis then have "x ! (length y) = \" using False by (simp add: nth_append) then have False using assms(2) False by (simp add: order_less_le) then show ?thesis by simp qed qed lemma numlistlist_inj: "numlistlist ns1 = numlistlist ns2 \ ns1 = ns2" proof (induction ns1 arbitrary: ns2) case Nil then show ?case using numlistlist_def by simp next case (Cons n ns1) have 1: "numlistlist (n # ns1) = (numlist n @ [\]) @ numlistlist ns1" using numlistlist_def by simp then have "numlistlist ns2 = (numlist n @ [\]) @ numlistlist ns1" using Cons by simp then have "ns2 \ []" using numlistlist_Nil by auto then have 2: "ns2 = hd ns2 # tl ns2" using `ns2 \ []` by simp then have 3: "numlistlist ns2 = (numlist (hd ns2) @ [\]) @ numlistlist (tl ns2)" using numlistlist_def by (metis concat.simps(2) list.simps(9)) have 4: "hd ns2 = n" using symbols_lt_prefix_eq 1 3 symbols_lt_numlist numlist_inj Cons by metis then have "numlistlist ns2 = (numlist n @ [\]) @ numlistlist (tl ns2)" using 3 by simp then have "numlistlist ns1 = numlistlist (tl ns2)" using 1 by (simp add: Cons.prems) then have "ns1 = tl ns2" using Cons by simp then show ?case using 2 4 by simp qed lemma numlistlist_append: "numlistlist (xs @ ys) = numlistlist xs @ numlistlist ys" using numlistlist_def by simp text \ Similar to @{text "\_\\<^sub>N"} and @{text "\_\\<^sub>N\<^sub>L"}, the tape contents for a list of list of numbers: \ definition nllcontents :: "nat list list \ (nat \ symbol)" ("\_\\<^sub>N\<^sub>L\<^sub>L") where "\nss\\<^sub>N\<^sub>L\<^sub>L \ \numlistlist nss\" lemma clean_tape_nllcontents: "clean_tape (\ns\\<^sub>N\<^sub>L\<^sub>L, i)" by (simp add: nllcontents_def proper_symbols_numlistlist) lemma nllcontents_Nil: "\[]\\<^sub>N\<^sub>L\<^sub>L = \[]\" using nllcontents_def by (simp add: numlistlist_Nil) text \ Similar to @{const nlength} and @{const nllength}, the length of the representation of a list of lists of numbers: \ definition nlllength :: "nat list list \ nat" where "nlllength nss \ length (numlistlist nss)" lemma nlllength: "nlllength nss = (\ns\nss. Suc (nllength ns))" using nlllength_def numlistlist_def nllength_def by (induction nss) simp_all lemma nlllength_Nil [simp]: "nlllength [] = 0" using nlllength_def numlistlist_def by simp lemma nlllength_Cons: "nlllength (ns # nss) = Suc (nllength ns) + nlllength nss" by (simp add: nlllength) lemma length_le_nlllength: "length nss \ nlllength nss" using nlllength sum_list_mono[of nss "\_. 1" "\ns. Suc (nllength ns)"] sum_list_const[of 1 nss] by simp lemma member_le_nlllength_1: "ns \ set nss \ nllength ns \ nlllength nss - 1" using nlllength by (induction nss) auto lemma nlllength_take: assumes "i < length nss" shows "nlllength (take i nss) + nllength (nss ! i) < nlllength nss" proof - have "nss = take i nss @ [nss ! i] @ drop (Suc i) nss" using assms by (metis Cons_eq_appendI append_self_conv2 id_take_nth_drop) then have "numlistlist nss = numlistlist (take i nss) @ numlistlist [nss ! i] @ numlistlist (drop (Suc i) nss)" using numlistlist_append by metis then have "nlllength nss = nlllength (take i nss) + nlllength [nss ! i] + nlllength (drop (Suc i) nss)" by (simp add: nlllength_def) then have "nlllength nss \ nlllength (take i nss) + nlllength [nss ! i]" by simp then have "nlllength nss \ nlllength (take i nss) + Suc (nllength (nss ! i))" using nlllength by simp then show ?thesis by simp qed lemma take_drop_numlistlist: assumes "i < length ns" shows "take (Suc (nllength (ns ! i))) (drop (nlllength (take i ns)) (numlistlist ns)) = numlist (ns ! i) @ [\]" proof - have "numlistlist ns = numlistlist (take i ns) @ numlistlist (drop i ns)" using numlistlist_append by (metis append_take_drop_id) moreover have "numlistlist (drop i ns) = numlistlist [ns ! i] @ numlistlist (drop (Suc i) ns)" using assms numlistlist_append by (metis Cons_nth_drop_Suc append_Cons self_append_conv2) ultimately have "numlistlist ns = numlistlist (take i ns) @ numlistlist [ns ! i] @ numlistlist (drop (Suc i) ns)" by simp then have "drop (nlllength (take i ns)) (numlistlist ns) = numlistlist [ns ! i] @ numlistlist (drop (Suc i) ns)" by (simp add: nlllength_def) then have "drop (nlllength (take i ns)) (numlistlist ns) = numlist (ns ! i) @ [\] @ numlistlist (drop (Suc i) ns)" using numlistlist_def by simp then show ?thesis by (simp add: nllength_def) qed corollary take_drop_numlistlist': assumes "i < length ns" shows "take (nllength (ns ! i)) (drop (nlllength (take i ns)) (numlistlist ns)) = numlist (ns ! i)" using take_drop_numlistlist[OF assms] nllength_def by (metis append_assoc append_eq_conv_conj append_take_drop_id) corollary numlistlist_take_at_term: assumes "i < length ns" shows "numlistlist ns ! (nlllength (take i ns) + nllength (ns ! i)) = \" using assms take_drop_numlistlist nlllength_def numlistlist_append - by (smt (z3) append_eq_conv_conj append_take_drop_id lessI nllength_def nth_append_length nth_append_length_plus nth_take) + by (smt (verit) append_eq_conv_conj append_take_drop_id lessI nllength_def nth_append_length nth_append_length_plus nth_take) lemma nlllength_take_Suc: assumes "i < length ns" shows "nlllength (take i ns) + Suc (nllength (ns ! i)) = nlllength (take (Suc i) ns)" proof - have "ns = take i ns @ [ns ! i] @ drop (Suc i) ns" using assms by (metis Cons_eq_appendI append_self_conv2 id_take_nth_drop) then have "numlistlist ns = numlistlist (take i ns) @ numlistlist [ns ! i] @ numlistlist (drop (Suc i) ns)" using numlistlist_append by metis then have "nlllength ns = nlllength (take i ns) + nlllength [ns ! i] + nlllength (drop (Suc i) ns)" by (simp add: nlllength_def) moreover have "nlllength ns = nlllength (take (Suc i) ns) + nlllength (drop (Suc i) ns)" using numlistlist_append by (metis append_take_drop_id length_append nlllength_def) ultimately have "nlllength (take (Suc i) ns) = nlllength (take i ns) + nlllength [ns ! i]" by simp then show ?thesis using nlllength by simp qed lemma numlistlist_take_at: assumes "i < length ns" and "j < nllength (ns ! i)" shows "numlistlist ns ! (nlllength (take i ns) + j) = numlist (ns ! i) ! j" proof - have "ns = take i ns @ [ns ! i] @ drop (Suc i) ns" using assms by (metis Cons_eq_appendI append_self_conv2 id_take_nth_drop) then have "numlistlist ns = (numlistlist (take i ns) @ numlistlist [ns ! i]) @ numlistlist (drop (Suc i) ns)" using numlistlist_append by (metis append_assoc) moreover have "nlllength (take i ns) + j < nlllength (take i ns) + nlllength [ns ! i]" using assms(2) nlllength by simp ultimately have "numlistlist ns ! (nlllength (take i ns) + j) = (numlistlist (take i ns) @ numlistlist [ns ! i]) ! (nlllength (take i ns) + j)" by (metis length_append nlllength_def nth_append) also have "... = numlistlist [ns ! i] ! j" by (simp add: nlllength_def) also have "... = (numlist (ns ! i) @ [\]) ! j" using numlistlist_def by simp also have "... = numlist (ns ! i) ! j" using assms(2) by (metis nllength_def nth_append) finally show ?thesis . qed lemma nllcontents_rneigh_5: assumes "i < length ns" shows "rneigh (\ns\\<^sub>N\<^sub>L\<^sub>L, Suc (nlllength (take i ns))) {\} (nllength (ns ! i))" proof (rule rneighI) let ?tp = "(\ns\\<^sub>N\<^sub>L\<^sub>L, Suc (nlllength (take i ns)))" show "fst ?tp (snd ?tp + nllength (ns ! i)) \ {\}" proof - have "snd ?tp + nllength (ns ! i) \ nlllength ns" using nlllength_take assms by (simp add: Suc_leI) then have "fst ?tp (snd ?tp + nllength (ns ! i)) = numlistlist ns ! (nlllength (take i ns) + nllength (ns ! i))" using nllcontents_def contents_inbounds nlllength_def by simp then show ?thesis using numlistlist_take_at_term assms by simp qed show "fst ?tp (snd ?tp + j) \ {\}" if "j < nllength (ns ! i)" for j proof - have "snd ?tp + nllength (ns ! i) \ nlllength ns" using nlllength_take assms by (simp add: Suc_leI) then have "snd ?tp + j \ nlllength ns" using that by simp then have "fst ?tp (snd ?tp + j) = numlistlist ns ! (nlllength (take i ns) + j)" using nllcontents_def contents_inbounds nlllength_def by simp then have "fst ?tp (snd ?tp + j) = numlist (ns ! i) ! j" using assms that numlistlist_take_at by simp moreover have "numlist (ns ! i) ! j \ \" using numlist_234 that nllength_def nth_mem by fastforce ultimately show ?thesis by simp qed qed text \ A tape containing a list of lists of numbers, with the tape head after all the symbols: \ abbreviation nlltape :: "nat list list \ tape" where "nlltape ns \ (\ns\\<^sub>N\<^sub>L\<^sub>L, Suc (nlllength ns))" text \ Like before but with the tape head or at the beginning of the $i$-th list element: \ abbreviation nlltape' :: "nat list list \ nat \ tape" where "nlltape' ns i \ (\ns\\<^sub>N\<^sub>L\<^sub>L, Suc (nlllength (take i ns)))" lemma nlltape'_0: "nlltape' ns 0 = (\ns\\<^sub>N\<^sub>L\<^sub>L, 1)" using nlllength by simp lemma nlltape'_tape_read: "|.| (nlltape' nss i) = 0 \ i \ length nss" proof - have "|.| (nlltape' nss i) = 0" if "i \ length nss" for i proof - have "nlltape' nss i \ (\nss\\<^sub>N\<^sub>L\<^sub>L, Suc (nlllength nss))" using that by simp then show ?thesis using nllcontents_def contents_outofbounds nlllength_def by (metis Suc_eq_plus1 add.left_neutral fst_conv less_Suc0 less_add_eq_less snd_conv) qed moreover have "|.| (nlltape' nss i) \ 0" if "i < length nss" for i using that nllcontents_def contents_inbounds nlllength_def nlllength_take proper_symbols_numlistlist by (metis Suc_leI add_Suc_right diff_Suc_1 fst_conv less_add_same_cancel1 less_le_trans not_add_less2 plus_1_eq_Suc snd_conv zero_less_Suc) ultimately show ?thesis by (meson le_less_linear) qed subsection \Appending an element\ text \ The next Turing machine is very similar to @{const tm_append}, just with terminator symbol @{text \} instead of @{text "\"}. It appends a list of numbers given on tape $j_2$ to the list of lists of numbers on tape $j_1$. \ definition tm_appendl :: "tapeidx \ tapeidx \ machine" where "tm_appendl j1 j2 \ tm_right_until j1 {\} ;; tm_cp_until j2 j1 {\} ;; tm_cr j2 ;; tm_char j1 \" lemma tm_appendl_tm: assumes "0 < j1" and "G \ 6" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_appendl j1 j2)" unfolding tm_appendl_def using assms tm_right_until_tm tm_cp_until_tm tm_char_tm tm_cr_tm by simp locale turing_machine_appendl = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_right_until j1 {\}" definition "tm2 \ tm1 ;; tm_cp_until j2 j1 {\}" definition "tm3 \ tm2 ;; tm_cr j2" definition "tm4 \ tm3 ;; tm_char j1 \" lemma tm4_eq_tm_append: "tm4 = tm_appendl j1 j2" unfolding tm4_def tm3_def tm2_def tm1_def tm_appendl_def by simp context fixes tps0 :: "tape list" and k i1 :: nat and ns :: "nat list" and nss :: "nat list list" assumes jk: "length tps0 = k" "j1 < k" "j2 < k" "j1 \ j2" "0 < j1" and i1: "i1 \ Suc (nlllength nss)" assumes tps0: "tps0 ! j1 = (\nss\\<^sub>N\<^sub>L\<^sub>L, i1)" "tps0 ! j2 = (\ns\\<^sub>N\<^sub>L, 1)" begin definition "tps1 \ tps0[j1 := nlltape nss]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (Suc (nlllength nss) - i1)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: jk) let ?l = "Suc (nlllength nss) - i1" show "rneigh (tps0 ! j1) {0} ?l" proof (rule rneighI) show "(tps0 ::: j1) (tps0 :#: j1 + ?l) \ {0}" using tps0 jk nllcontents_def nlllength_def by simp show "(tps0 ::: j1) (tps0 :#: j1 + i) \ {0}" if "i < Suc (nlllength nss) - i1" for i proof - have "i1 + i < Suc (nlllength nss)" using that i1 by simp then show ?thesis using proper_symbols_numlistlist nlllength_def tps0 nllcontents_def contents_def by (metis One_nat_def Suc_leI diff_Suc_1 fst_conv less_Suc_eq_0_disj less_nat_zero_code singletonD snd_conv) qed qed show "ttt = Suc (Suc (nlllength nss) - i1)" using assms(1) . show "tps1 = tps0[j1 := tps0 ! j1 |+| Suc (nlllength nss) - i1]" using tps1_def tps0 i1 by simp qed definition "tps2 \ tps0 [j1 := (\numlistlist nss @ numlist ns\, Suc (nlllength nss) + nllength ns), j2 := (\ns\\<^sub>N\<^sub>L, Suc (nllength ns))]" lemma tm2 [transforms_intros]: assumes "ttt = 3 + nlllength nss - i1 + nllength ns" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: jk tps1_def) have j1: "tps1 ! j1 = nlltape nss" using tps1_def by (simp add: jk) have j2: "tps1 ! j2 = (\ns\\<^sub>N\<^sub>L, 1)" using tps1_def tps0 jk by simp show "rneigh (tps1 ! j2) {0} (nllength ns)" proof (rule rneighI) show "(tps1 ::: j2) (tps1 :#: j2 + nllength ns) \ {0}" using j2 by (simp add: nlcontents_def nllength_def) show "\i. i < nllength ns \ (tps1 ::: j2) (tps1 :#: j2 + i) \ {0}" using j2 tps0 contents_def nlcontents_def nllength_def proper_symbols_numlist by fastforce qed show "tps2 = tps1 [j2 := tps1 ! j2 |+| nllength ns, j1 := implant (tps1 ! j2) (tps1 ! j1) (nllength ns)]" (is "_ = ?rhs") proof - have "implant (tps1 ! j2) (tps1 ! j1) (nllength ns) = implant (\ns\\<^sub>N\<^sub>L, 1) (nlltape nss) (nllength ns)" using j1 j2 by simp also have "... = (\numlistlist nss @ (take (nllength ns) (drop (1 - 1) (numlist ns)))\, Suc (length (numlistlist nss)) + nllength ns)" using implant_contents nlcontents_def nllength_def nllcontents_def nlllength_def by simp also have "... = (\numlistlist nss @ numlist ns\, Suc (length (numlistlist nss)) + nllength ns)" by (simp add: nllength_def) also have "... = (\numlistlist nss @ numlist ns\, Suc (nlllength nss) + nllength ns)" using nlllength_def by simp also have "... = tps2 ! j1" using jk tps2_def by (metis nth_list_update_eq nth_list_update_neq) finally have "implant (tps1 ! j2) (tps1 ! j1) (nllength ns) = tps2 ! j1" . then have "tps2 ! j1 = implant (tps1 ! j2) (tps1 ! j1) (nllength ns)" by simp then have "tps2 ! j1 = ?rhs ! j1" using tps1_def jk by (metis length_list_update nth_list_update_eq) moreover have "tps2 ! j2 = ?rhs ! j2" using tps2_def tps1_def jk j2 by simp moreover have "tps2 ! j = ?rhs ! j" if "j < length tps2" "j \ j1" "j \ j2" for j using that tps2_def tps1_def by simp moreover have "length tps2 = length ?rhs" using tps1_def tps2_def by simp ultimately show ?thesis using nth_equalityI by blast qed show "ttt = Suc (Suc (nlllength nss) - i1) + Suc (nllength ns)" using assms(1) j1 tps0 i1 by simp qed definition "tps3 \ tps0 [j1 := (\numlistlist nss @ numlist ns\, Suc (nlllength nss) + nllength ns)]" lemma tm3 [transforms_intros]: assumes "ttt = 6 + nlllength nss - i1 + 2 * nllength ns" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: jk i1 tps2_def) let ?tp1 = "(\numlistlist nss @ numlist ns\, Suc (nlllength nss + nllength ns))" let ?tp2 = "(\ns\\<^sub>N\<^sub>L, Suc (nllength ns))" show "clean_tape (tps2 ! j2)" using tps2_def jk by (simp add: clean_tape_nlcontents) show "tps3 = tps2[j2 := tps2 ! j2 |#=| 1]" using tps3_def tps2_def jk tps0(2) by (metis fst_eqD length_list_update list_update_overwrite list_update_same_conv nth_list_update) show "ttt = 3 + nlllength nss - i1 + nllength ns + (tps2 :#: j2 + 2)" using assms tps2_def jk tps0 i1 by simp qed definition "tps4 \ tps0 [j1 := (\numlistlist (nss @ [ns])\, Suc (nlllength (nss @ [ns])))]" lemma tm4: assumes "ttt = 7 + nlllength nss - i1 + 2 * nllength ns" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: jk tps3_def time: jk i1 assms) show "tps4 = tps3[j1 := tps3 ! j1 |:=| \ |+| 1]" (is "tps4 = ?tps") proof - have "tps3 ! j1 = (\numlistlist nss @ numlist ns\, Suc (nlllength nss) + nllength ns)" using tps3_def jk by simp then have "?tps = tps0[j1 := (\numlistlist nss @ numlist ns\, Suc (nlllength nss + nllength ns)) |:=| \ |+| 1]" using tps3_def by simp moreover have "(\numlistlist nss @ numlist ns\, Suc (nlllength nss + nllength ns)) |:=| \ |+| 1 = (\numlistlist (nss @ [ns])\, Suc (nlllength (nss @ [ns])))" (is "?lhs = ?rhs") proof - have "?lhs = (\numlistlist nss @ numlist ns\(Suc (nlllength nss + nllength ns) := \), Suc (Suc (nlllength nss + nllength ns)))" by simp also have "... = (\numlistlist nss @ numlist ns\(Suc (nlllength nss + nllength ns) := \), Suc (nlllength (nss @ [ns])))" using nlllength_def numlistlist_def nllength_def numlist_def by simp also have "... = (\(numlistlist nss @ numlist ns) @ [\]\, Suc (nlllength (nss @ [ns])))" using contents_snoc nlllength_def nllength_def by (metis length_append) also have "... = (\numlistlist nss @ numlist ns @ [\]\, Suc (nlllength (nss @ [ns])))" by simp also have "... = (\numlistlist (nss @ [ns])\, Suc (nlllength (nss @ [ns])))" using numlistlist_def by simp finally show "?lhs = ?rhs" . qed ultimately show ?thesis using tps4_def by auto qed qed end (* context *) end (* locale turing_machine_appendl *) lemma transforms_tm_appendlI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and ttt k i1 :: nat and ns :: "nat list" and nss :: "nat list list" assumes "0 < j1" assumes "length tps = k" "j1 < k" "j2 < k" "j1 \ j2" and "i1 \ Suc (nlllength nss)" assumes "tps ! j1 = (\nss\\<^sub>N\<^sub>L\<^sub>L, i1)" "tps ! j2 = (\ns\\<^sub>N\<^sub>L, 1)" assumes "ttt = 7 + nlllength nss - i1 + 2 * nllength ns" assumes "tps' = tps [j1 := nlltape (nss @ [ns])]" shows "transforms (tm_appendl j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_appendl j1 j2 . show ?thesis using loc.tps4_def loc.tm4 loc.tm4_eq_tm_append assms nllcontents_def by simp qed subsection \Extending with a list\ text \ The next Turing machine extends a list of lists of numbers with another list of lists of numbers. It is in fact the same TM as for extending a list of numbers because in both cases extending means simply copying the contents from one tape to another. We introduce a new name for this TM and express the semantics in terms of lists of lists of numbers. The proof is almost the same except with @{const nllength} replaced by @{const nlllength} and so on. \ definition tm_extendl :: "tapeidx \ tapeidx \ machine" where "tm_extendl \ tm_extend" lemma tm_extendl_tm: assumes "0 < j1" and "G \ 4" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_extendl j1 j2)" unfolding tm_extendl_def using assms tm_extend_tm by simp locale turing_machine_extendl = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_cp_until j2 j1 {\}" definition "tm2 \ tm1 ;; tm_cr j2" lemma tm2_eq_tm_extendl: "tm2 = tm_extendl j1 j2" unfolding tm2_def tm2_def tm1_def tm_extendl_def tm_extend_def by simp context fixes tps0 :: "tape list" and k :: nat and nss1 nss2 :: "nat list list" assumes jk: "0 < j1" "j1 < k" "j2 < k" "j1 \ j2" "length tps0 = k" assumes tps0: "tps0 ! j1 = nlltape nss1" "tps0 ! j2 = (\nss2\\<^sub>N\<^sub>L\<^sub>L, 1)" begin definition "tps1 \ tps0 [j1 := nlltape (nss1 @ nss2), j2 := nlltape nss2]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (nlllength nss2)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps1_def tps0 jk time: assms) let ?n = "nlllength nss2" show "rneigh (tps0 ! j2) {\} ?n" proof (rule rneighI) show "(tps0 ::: j2) (tps0 :#: j2 + nlllength nss2) \ {\}" using tps0 nllcontents_def nlllength_def jk by simp show "\i. i < nlllength nss2 \ (tps0 ::: j2) (tps0 :#: j2 + i) \ {\}" using tps0 jk contents_def nllcontents_def nlllength_def proper_symbols_numlistlist by fastforce qed show "tps1 = tps0 [j2 := tps0 ! j2 |+| nlllength nss2, j1 := implant (tps0 ! j2) (tps0 ! j1) (nlllength nss2)]" proof - have "implant (\nss2\\<^sub>N\<^sub>L\<^sub>L, 1) (nlltape nss1) (nlllength nss2) = nlltape (nss1 @ nss2)" using nllcontents_def nlllength_def implant_contents by (simp add: numlistlist_append) moreover have "tps1 ! j2 = tps0 ! j2 |+| nlllength nss2" using tps0 jk tps1_def by simp ultimately show ?thesis using tps0 jk tps1_def by (metis fst_conv list_update_swap plus_1_eq_Suc snd_conv) qed qed definition "tps2 \ tps0[j1 := nlltape (nss1 @ nss2)]" lemma tm2: assumes "ttt = 4 + 2 * nlllength nss2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps1_def tps2_def jk time: assms tps1_def jk) show "clean_tape (tps1 ! j2)" using tps1_def jk clean_tape_nllcontents by simp show "tps2 = tps1[j2 := tps1 ! j2 |#=| 1]" using tps1_def jk tps2_def tps0(2) by (metis fst_conv length_list_update list_update_id list_update_overwrite nth_list_update) qed end (* context tps0 *) end (* locale turing_machine_extendl *) lemma transforms_tm_extendlI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k :: nat and nss1 nss2 :: "nat list list" assumes "0 < j1" "j1 < k" "j2 < k" "j1 \ j2" "length tps = k" assumes "tps ! j1 = nlltape nss1" "tps ! j2 = (\nss2\\<^sub>N\<^sub>L\<^sub>L, 1)" assumes "ttt = 4 + 2 * nlllength nss2" assumes "tps' = tps[j1 := nlltape (nss1 @ nss2)]" shows "transforms (tm_extendl j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_extendl j1 j2 . show ?thesis using loc.tm2_eq_tm_extendl loc.tm2 loc.tps2_def assms by simp qed text \ The next Turing machine erases the appended list. \ definition tm_extendl_erase :: "tapeidx \ tapeidx \ machine" where "tm_extendl_erase j1 j2 \ tm_extendl j1 j2 ;; tm_erase_cr j2" lemma tm_extendl_erase_tm: assumes "0 < j1" and "0 < j2" and "G \ 4" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_extendl_erase j1 j2)" unfolding tm_extendl_erase_def using assms tm_extendl_tm tm_erase_cr_tm by simp lemma transforms_tm_extendl_eraseI [transforms_intros]: fixes tps tps' :: "tape list" and j1 j2 :: tapeidx and ttt k :: nat and nss1 nss2 :: "nat list list" assumes "0 < j1" "j1 < k" "j2 < k" "j1 \ j2" "0 < j2" "length tps = k" assumes "tps ! j1 = nlltape nss1" "tps ! j2 = (\nss2\\<^sub>N\<^sub>L\<^sub>L, 1)" assumes "ttt = 11 + 4 * nlllength nss2 " assumes "tps' = tps [j1 := nlltape (nss1 @ nss2), j2 := (\[]\, 1)]" shows "transforms (tm_extendl_erase j1 j2) tps ttt tps'" unfolding tm_extendl_erase_def proof (tform tps: assms) let ?zs = "numlistlist nss2" show "tps[j1 := nlltape (nss1 @ nss2)] ::: j2 = \?zs\" using assms by (simp add: nllcontents_def) show "proper_symbols ?zs" using proper_symbols_numlistlist by simp show "ttt = 4 + 2 * nlllength nss2 + (tps[j1 := nlltape (nss1 @ nss2)] :#: j2 + 2 * length (numlistlist nss2) + 6)" using assms nlllength_def by simp qed subsection \Moving to the next element\ text \ Iterating over a list of lists of numbers works with the same Turing machine, @{const tm_nextract}, as for lists of numbers. We just have to set the parameter $z$ to the terminator symbol @{text \}. For the proof we can also just copy from the previous section, replacing the symbol @{text "\"} by @{text \} and @{const nllength} by @{const nlllength}, etc. \null \ locale turing_machine_nextract_5 = fixes j1 j2 :: tapeidx begin definition "tm1 \ tm_erase_cr j2" definition "tm2 \ tm1 ;; tm_cp_until j1 j2 {\}" definition "tm3 \ tm2 ;; tm_cr j2" definition "tm4 \ tm3 ;; tm_right j1" lemma tm4_eq_tm_nextract: "tm4 = tm_nextract \ j1 j2" unfolding tm1_def tm2_def tm3_def tm4_def tm_nextract_def by simp context fixes tps0 :: "tape list" and k idx :: nat and nss :: "nat list list" and dummy :: "nat list" assumes jk: "j1 < k" "j2 < k" "0 < j1" "0 < j2" "j1 \ j2" "length tps0 = k" and idx: "idx < length nss" and tps0: "tps0 ! j1 = nlltape' nss idx" "tps0 ! j2 = (\dummy\\<^sub>N\<^sub>L, 1)" begin definition "tps1 \ tps0[j2 := (\[]\\<^sub>N\<^sub>L, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 7 + 2 * nllength dummy" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: jk idx tps0 tps1_def assms) show "proper_symbols (numlist dummy)" using proper_symbols_numlist by simp show "tps1 = tps0[j2 := (\[]\, 1)]" using tps1_def by (simp add: nlcontents_def numlist_Nil) show "tps0 ::: j2 = \numlist dummy\" using tps0 by (simp add: nlcontents_def) show "ttt = tps0 :#: j2 + 2 * length (numlist dummy) + 6" using tps0 assms nllength_def by simp qed definition "tps2 \ tps0 [j1 := (\nss\\<^sub>N\<^sub>L\<^sub>L, nlllength (take (Suc idx) nss)), j2 := (\nss ! idx\\<^sub>N\<^sub>L, Suc (nllength (nss ! idx)))]" lemma tm2 [transforms_intros]: assumes "ttt = 7 + 2 * nllength dummy + Suc (nllength (nss ! idx))" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: jk idx tps0 tps2_def tps1_def time: assms) show "rneigh (tps1 ! j1) {\} (nllength (nss ! idx))" using tps1_def tps0 jk by (simp add: idx nllcontents_rneigh_5) show "tps2 = tps1 [j1 := tps1 ! j1 |+| nllength (nss ! idx), j2 := implant (tps1 ! j1) (tps1 ! j2) (nllength (nss ! idx))]" (is "?l = ?r") proof (rule nth_equalityI) show len: "length ?l = length ?r" using tps1_def tps2_def jk by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - consider "i = j1" | "i = j2" | "i \ j1 \ i \ j2" by auto then show ?thesis proof (cases) case 1 then show ?thesis using tps0 that tps1_def tps2_def jk nlllength_take_Suc[OF idx] idx by simp next case 2 then have lhs: "?l ! i = (\nss ! idx\\<^sub>N\<^sub>L, Suc (nllength (nss ! idx)))" using tps2_def jk by simp let ?i = "Suc (nlllength (take idx nss))" have i1: "?i > 0" by simp have "nllength (nss ! idx) + (?i - 1) \ nlllength nss" using idx nlllength_take by (metis add.commute diff_Suc_1 less_or_eq_imp_le) then have i2: "nllength (nss ! idx) + (?i - 1) \ length (numlistlist nss)" using nlllength_def by simp have "?r ! i = implant (tps1 ! j1) (tps1 ! j2) (nllength (nss ! idx))" using 2 tps1_def jk by simp also have "... = implant (\nss\\<^sub>N\<^sub>L\<^sub>L, ?i) (\[]\\<^sub>N\<^sub>L, 1) (nllength (nss ! idx))" using tps1_def jk tps0 by simp also have "... = (\[] @ (take (nllength (nss ! idx)) (drop (?i - 1) (numlistlist nss)))\, Suc (length []) + nllength (nss ! idx))" using implant_contents[OF i1 i2] nllcontents_def nlcontents_def numlist_Nil by (metis One_nat_def list.size(3)) finally have "?r ! i = (\[] @ (take (nllength (nss ! idx)) (drop (?i - 1) (numlistlist nss)))\, Suc (length []) + nllength (nss ! idx))" . then have "?r ! i = (\take (nllength (nss ! idx)) (drop (nlllength (take idx nss)) (numlistlist nss))\, Suc (nllength (nss ! idx)))" by simp then have "?r ! i = (\numlist (nss ! idx)\, Suc (nllength (nss ! idx)))" using take_drop_numlistlist'[OF idx] by simp then show ?thesis using lhs nlcontents_def by simp next case 3 then show ?thesis using that tps1_def tps2_def jk by simp qed qed qed qed definition "tps3 \ tps0 [j1 := (\nss\\<^sub>N\<^sub>L\<^sub>L, nlllength (take (Suc idx) nss)), j2 := (\nss ! idx\\<^sub>N\<^sub>L, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 11 + 2 * nllength dummy + 2 * (nllength (nss ! idx))" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: jk idx tps0 tps2_def tps3_def assms) show "clean_tape (tps2 ! j2)" using tps2_def jk clean_tape_nlcontents by simp qed definition "tps4 \ tps0 [j1 := nlltape' nss (Suc idx), j2 := (\nss ! idx\\<^sub>N\<^sub>L, 1)]" lemma tm4: assumes "ttt = 12 + 2 * nllength dummy + 2 * (nllength (nss ! idx))" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def by (tform tps: jk idx tps0 tps3_def tps4_def assms) end (* context *) end (* locale turing_machine_nextract *) lemma transforms_tm_nextract_5I [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k idx :: nat and nss :: "nat list list" and dummy :: "nat list" assumes "j1 < k" "j2 < k" "0 < j1" "0 < j2" "j1 \ j2" "length tps = k" and "idx < length nss" assumes "tps ! j1 = nlltape' nss idx" "tps ! j2 = (\dummy\\<^sub>N\<^sub>L, 1)" assumes "ttt = 12 + 2 * nllength dummy + 2 * (nllength (nss ! idx))" assumes "tps' = tps [j1 := nlltape' nss (Suc idx), j2 := (\nss ! idx\\<^sub>N\<^sub>L, 1)]" shows "transforms (tm_nextract \ j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_nextract_5 j1 j2 . show ?thesis using assms loc.tm4 loc.tps4_def loc.tm4_eq_tm_nextract by simp qed end diff --git a/thys/Cook_Levin/NP.thy b/thys/Cook_Levin/NP.thy --- a/thys/Cook_Levin/NP.thy +++ b/thys/Cook_Levin/NP.thy @@ -1,702 +1,702 @@ chapter \Time complexity\label{s:TC}\ theory NP imports Elementary Composing Symbol_Ops begin text \ In order to formulate the Cook-Levin theorem we need to formalize \SAT{} and $\NP$-completeness. This chapter is devoted to the latter and hence introduces the complexity class $\NP$ and the concept of polynomial-time many-one reduction. Moreover, although not required for the Cook-Levin theorem, it introduces the class $\mathcal{P}$ and contains a proof of $\mathcal{P} \subseteq \NP$ (see Section~\ref{s:TC-subseteq}). The chapter concludes with some easy results about $\mathcal{P}$, $\NP$ and reducibility in Section~\ref{s:TC-more}. \ section \The time complexity classes DTIME, $\mathcal{P}$, and $\NP$\label{s:TC-NP}\ text \ Arora and Barak~\cite[Definitions 1.12, 1.13]{ccama} define $\mathrm{DTIME}(T(n))$ as the set of all languages that can be decided in time $c \cdot T(n)$ for some $c > 0$ and $\mathcal{P} = \bigcup_{c\geq1}\mathrm{DTIME}(n^c)$. However since $0^c = 0$ for $c\geq 1$, this means that for a language $L$ to be in $\mathcal{P}$, the Turing machine deciding $L$ must check the empty string in zero steps. For a Turing machine to halt in zero steps, it must start in the halting state, which limits its usefulness. Because of this technical issue we define $\mathrm{DTIME}(T(n))$ as the set of all languages that can be decided with a running time in $O(T(n))$, which seems a common enough alternative~\cite{lv-aikc,sipser2006,cz-dtime}. \ text \ Languages are sets of strings, and deciding a language means computing its characteristic function. \ type_synonym language = "string set" definition characteristic :: "language \ (string \ string)" where "characteristic L \ (\x. [x \ L])" definition DTIME :: "(nat \ nat) \ language set" where "DTIME T \ {L. \k G M T'. turing_machine k G M \ big_oh T' T \ computes_in_time k M (characteristic L) T'}" definition complexity_class_P :: "language set" ("\

") where "\

\ \c\{1..}. DTIME (\n. n ^ c)" text \ A language $L$ is in $\NP$ if there is a polynomial $p$ and a polynomial-time Turing machine, called the \emph{verifier}, such that for all strings $x\in\bbOI^*$, \[ x\in L \longleftrightarrow \exists u\in\bbOI^{p(|x|)}: M(\langle x, u\rangle) = \bbbI. \] The string $u$ does not seem to have a name in general, but in case the verifier outputs $\bbbI$ on input $\langle x, u\rangle$ it is called a \emph{certificate} for $x$~\cite[Definition~2.1]{ccama}. \ definition complexity_class_NP :: "language set" ("\\

") where "\\

\ {L. \k G M p T fverify. turing_machine k G M \ polynomial p \ big_oh_poly T \ computes_in_time k M fverify T \ (\x. x \ L \ (\u. length u = p (length x) \ fverify \x, u\ = [\]))}" text \ The definition of $\NP$ is the one place where we need an actual polynomial function, namely $p$, rather than a function that is merely bounded by a polynomial. This raises the question as to the definition of a polynomial function. Arora and Barak~\cite{ccama} do not seem to give a definition in the context of $\NP$ but explicitly state that polynomial functions are mappings $\nat \to \nat$. Presumably they also have the form $f(n) = \sum_{i=0}^d a_i\cdot n^i$, as polynomials do. We have filled in the gap in this definition in Section~\ref{s:tm-basic-bigoh} by letting the coefficients $a_i$ be natural numbers. Regardless of whether this is the meaning intended by Arora and Barak, the choice is justified because with it the verifier-based definition of $\NP$ is equivalent to the original definition via nondeterministic Turing machines (NTMs). In the usual equivalence proof (for example, Arora and Barak~\cite[Theorem~2.6]{ccama}) a verifier TM and an NTM are constructed. For the one direction, if a language is decided by a polynomial-time NTM then a verifier TM can be constructed that simulates the NTM on input $x$ by using the bits in the string $u$ for the nondeterministic choices. The strings $u$ have the length $p(|x|)$. So for this construction to work, there must be a polynomial $p$ that bounds the running time of the NTM. Clearly, a polynomial function with natural coefficients exists with that property. For the other direction, assume a language has a verifier TM where $p$ is a polynomial function with natural coefficients. An NTM deciding this language receives $x$ as input, then ``guesses'' a string $u$ of length $p(|x|)$, and then runs the verifier on the pair $\langle x, u\rangle$. For this NTM to run in polynomial time, $p$ must be computable in time polynomial in $|x|$. We have shown this to be the case in lemma @{thm [source] transforms_tm_polynomialI} in Section~\ref{s:tm-arithmetic-poly}. \ text \ A language $L_1$ is polynomial-time many-one reducible to a language $L_2$ if there is a polynomial-time computable function $f_\mathit{reduce}$ such that for all $x$, $x\in L_1$ iff.\ $f_\mathit{reduce}(x) \in L_2$. \ definition reducible (infix "\\<^sub>p" 50) where "L\<^sub>1 \\<^sub>p L\<^sub>2 \ \k G M T freduce. turing_machine k G M \ big_oh_poly T \ computes_in_time k M freduce T \ (\x. x \ L\<^sub>1 \ freduce x \ L\<^sub>2)" abbreviation NP_hard :: "language \ bool" where "NP_hard L \ \L'\\\

. L' \\<^sub>p L" definition NP_complete :: "language \ bool" where "NP_complete L \ L \ \\

\ NP_hard L" text \ Requiring $c \geq 1$ in the definition of $\mathcal{P}$ is not essential: \ lemma in_P_iff: "L \ \

\ (\c. L \ DTIME (\n. n ^ c))" proof assume "L \ \

" then show "\c. L \ DTIME (\n. n ^ c)" unfolding complexity_class_P_def using Suc_le_eq by auto next assume "\c. L \ DTIME (\n. n ^ c)" then obtain c where "L \ DTIME (\n. n ^ c)" by auto then obtain k G M T where M: "turing_machine k G M" "big_oh T (\n. n ^ c)" "computes_in_time k M (characteristic L) T" using DTIME_def by auto moreover have "big_oh T (\n. n ^ Suc c)" proof - obtain c0 n0 :: nat where c0n0: "\n>n0. T n \ c0 * n ^ c" using M(2) big_oh_def by auto have "\n>n0. T n \ c0 * n ^ Suc c" proof standard+ fix n assume "n0 < n" then have "n ^ c \ n ^ Suc c" using pow_mono by simp then show "T n \ c0 * n ^ Suc c" using c0n0 by (metis \n0 < n\ add.commute le_Suc_ex mult_le_mono2 trans_le_add2) qed then show ?thesis using big_oh_def by auto qed ultimately have "\c>0. L \ DTIME (\n. n ^ c)" using DTIME_def by blast then show "L \ \

" unfolding complexity_class_P_def by auto qed section \Restricting verifiers to one-bit output\label{s:np-alt}\ text \ The verifier Turing machine in the definition of $\NP$ can output any symbol sequence. In this section we restrict it to outputting only the symbol sequence \textbf{1} or \textbf{0}. This is equivalent to the definition because it is easy to check if a symbol sequence differs from \textbf{1} and if so change it to \textbf{0}, as we show below. The advantage of this restriction is that if we can make the TM halt with the output tape head on cell number~1, the output tape symbol read in the final step equals the output of the TM. We will exploit this in Chapter~\ref{s:Reducing}, where we show how to reduce any language $L\in\NP$ to \SAT{}. The next Turing machine checks if the symbol sequence on tape $j$ differs from the one-symbol sequence \textbf{1} and if so turns it into \textbf{0}. It thus ensures that the tape contains only one bit symbol. \null \ definition tm_make_bit :: "tapeidx \ machine" where "tm_make_bit j \ tm_cr j ;; IF \rs. rs ! j = \ THEN tm_right j ;; IF \rs. rs ! j = \ THEN [] ELSE tm_set j [\] ENDIF ELSE tm_set j [\] ENDIF" lemma tm_make_bit_tm: assumes "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_make_bit j)" unfolding tm_make_bit_def using assms tm_right_tm tm_set_tm tm_cr_tm Nil_tm turing_machine_branch_turing_machine by simp locale turing_machine_make_bit = fixes j :: tapeidx begin definition "tm1 \ tm_cr j" definition "tmT1 \ tm_right j" definition "tmT12 \ IF \rs. rs ! j = \ THEN [] ELSE tm_set j [\] ENDIF" definition "tmT2 \ tmT1 ;; tmT12" definition "tm12 \ IF \rs. rs ! j = \ THEN tmT2 ELSE tm_set j [\] ENDIF" definition "tm2 \ tm1 ;; tm12" lemma tm2_eq_tm_make_bit: "tm2 \ tm_make_bit j" unfolding tm_make_bit_def tm2_def tm12_def tmT2_def tmT12_def tmT1_def tm1_def by simp context fixes tps0 :: "tape list" and zs :: "symbol list" assumes jk: "j < length tps0" and zs: "proper_symbols zs" assumes tps0: "tps0 ::: j = \zs\" begin lemma clean: "clean_tape (tps0 ! j)" using zs tps0 contents_clean_tape' by simp definition "tps1 \ tps0[j := (\zs\, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = tps0 :#: j + 2" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def by (tform tps: assms jk clean tps0 tps1_def) definition "tpsT1 \ tps0[j := (\zs\, 2)]" lemma tmT1 [transforms_intros]: assumes "ttt = 1" shows "transforms tmT1 tps1 ttt tpsT1" unfolding tmT1_def proof (tform tps: assms tps1_def jk) show "tpsT1 = tps1[j := tps1 ! j |+| 1] " using tps1_def tpsT1_def jk by (metis Suc_1 fst_conv list_update_overwrite nth_list_update_eq plus_1_eq_Suc snd_conv) qed definition "tpsT2 \ tps0 [j := if length zs \ 1 then (\zs\, 2) else (\[\]\, 1)]" lemma tmT12 [transforms_intros]: assumes "ttt = 14 + 2 * length zs" shows "transforms tmT12 tpsT1 ttt tpsT2" unfolding tmT12_def proof (tform tps: assms tpsT1_def tps0 jk zs) show "8 + tpsT1 :#: j + 2 * length zs + Suc (2 * length [\]) + 1 \ ttt" using tpsT1_def jk assms by simp have "read tpsT1 ! j = \zs\ 2" using tpsT1_def jk tapes_at_read' by (metis fst_conv length_list_update nth_list_update_eq snd_conv) moreover have "\zs\ 2 = \ \ length zs \ 1" using zs contents_def by (metis Suc_1 diff_Suc_1 less_imp_le_nat linorder_le_less_linear not_less_eq_eq zero_neq_numeral) ultimately have "read tpsT1 ! j = \ \ length zs \ 1" by simp then show "read tpsT1 ! j \ \ \ tpsT2 = tpsT1[j := (\[\]\, 1)]" "read tpsT1 ! j = \ \ tpsT2 = tpsT1" using tpsT1_def tpsT2_def jk by simp_all qed lemma tmT2 [transforms_intros]: assumes "ttt = 15 + 2 * length zs" shows "transforms tmT2 tps1 ttt tpsT2" unfolding tmT2_def by (tform time: assms) definition "tps2 \ tps0 [j := if zs = [\] then (\zs\, 2) else (\[\]\, 1)]" lemma tm12 [transforms_intros]: assumes "ttt = 17 + 2 * length zs" shows "transforms tm12 tps1 ttt tps2" unfolding tm12_def proof (tform tps: assms tps0 jk zs tps1_def) have "read tps1 ! j = \zs\ 1" using tps1_def jk tapes_at_read' by (metis fst_conv length_list_update nth_list_update_eq snd_conv) then have *: "read tps1 ! j = \ \ \zs\ 1 = \" by simp show "read tps1 ! j \ \ \ tps2 = tps1[j := (\[\]\, 1)]" using * tps2_def tps1_def by auto show "tps2 = tpsT2" if "read tps1 ! j = \" proof (cases "zs = [\]") case True then show ?thesis using * tps2_def tpsT2_def by simp next case False then have "\zs\ 1 = \" using * that by simp then have "length zs > 1" using False contents_def contents_outofbounds by (metis One_nat_def Suc_length_conv diff_Suc_1 length_0_conv linorder_neqE_nat nth_Cons_0 zero_neq_numeral) then show ?thesis using * tps2_def tpsT2_def by auto qed show "8 + tps1 :#: j + 2 * length zs + Suc (2 * length [\]) + 1 \ ttt" using tps1_def jk assms by simp qed lemma tm2: assumes "ttt = 19 + 2 * length zs + tps0 :#: j" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: assms tps0 jk zs tps1_def) end end (* locale *) lemma transforms_tm_make_bitI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and zs :: "symbol list" and ttt :: nat assumes "j < length tps" and "proper_symbols zs" assumes "tps ::: j = \zs\" assumes "ttt = 19 + 2 * length zs + tps :#: j" assumes "tps' = tps [j := if zs = [\] then (\zs\, 2) else (\[\]\, 1)]" shows "transforms (tm_make_bit j) tps ttt tps'" proof - interpret loc: turing_machine_make_bit j . show ?thesis using assms loc.tps2_def loc.tm2 loc.tm2_eq_tm_make_bit by simp qed lemma output_length_le_time: assumes "turing_machine k G M" and "tps ::: 1 = \zs\" and "proper_symbols zs" and "transforms M (snd (start_config k xs)) t tps" shows "length zs \ t" proof - have 1: "execute M (start_config k xs) t = (length M, tps)" using assms transforms_def transits_def by (metis (no_types, lifting) execute_after_halting_ge fst_conv start_config_def sndI) moreover have "k > 1" using assms(1) turing_machine_def by simp ultimately have "((execute M (start_config k xs) t) <:> 1) i = \" if "i > t" for i using assms blank_after_time that by (meson zero_less_one) then have "(tps ::: 1) i = \" if "i > t" for i using 1 that by simp then have *: "\zs\ i = \" if "i > t" for i using assms(2) that by simp show ?thesis proof (rule ccontr) assume "\ length zs \ t" then have "length zs > t" by simp then have "\zs\ (Suc t) \ \" using contents_inbounds assms(3) contents_def proper_symbols_ne0 by simp then show False using * by simp qed qed text \ This is the alternative definition of $\NP$, which restricts the verifier to output only strings of length one: \ lemma NP_output_len_1: "\\

= {L. \k G M p T fverify. turing_machine k G M \ polynomial p \ big_oh_poly T \ computes_in_time k M fverify T \ (\y. length (fverify y) = 1) \ (\x. x \ L \ (\u. length u = p (length x) \ fverify \x, u\ = [\]))}" (is "_ = ?M") proof show "?M \ \\

" using complexity_class_NP_def by fast define Q where "Q = (\L k G M p T fverify. turing_machine k G M \ polynomial p \ big_oh_poly T \ computes_in_time k M fverify T \ (\x. (x \ L) = (\u. length u = p (length x) \ fverify \x, u\ = [\])))" have alt: "complexity_class_NP = {L::language. \k G M p T fverify. Q L k G M p T fverify}" unfolding complexity_class_NP_def Q_def by simp show "\\

\ ?M" proof fix L assume "L \ \\

" then obtain k G M p T fverify where "Q L k G M p T fverify" using alt by blast then have ex: "turing_machine k G M \ polynomial p \ big_oh_poly T \ computes_in_time k M fverify T \ (\x. (x \ L) = (\u. length u = p (length x) \ fverify \x, u\ = [\]))" using Q_def by simp have "k \ 2" and "G \ 4" using ex turing_machine_def by simp_all define M' where "M' = M ;; tm_make_bit 1" define f' where "f' = (\y. if fverify y = [\] then [\] else [\])" define T' where "T' = (\n. 19 + 4 * T n)" have "turing_machine k G M'" unfolding M'_def using \2 \ k\ \4 \ G\ ex tm_make_bit_tm by simp moreover have "polynomial p" using ex by simp moreover have "big_oh_poly T'" using T'_def ex big_oh_poly_const big_oh_poly_prod big_oh_poly_sum by simp moreover have "computes_in_time k M' f' T'" proof fix y let ?cfg = "start_config k (string_to_symbols y)" obtain tps where 1: "tps ::: 1 = string_to_contents (fverify y)" and 2: "transforms M (snd ?cfg) (T (length y)) tps" using ex computes_in_timeD by blast have len_tps: "length tps \ 2" - by (smt (z3) "2" \2 \ k\ ex execute_num_tapes start_config_length less_le_trans numeral_2_eq_2 + by (smt (verit) "2" \2 \ k\ ex execute_num_tapes start_config_length less_le_trans numeral_2_eq_2 prod.sel(2) transforms_def transits_def zero_less_Suc) define zs where "zs = string_to_symbols (fverify y)" then have zs: "tps ::: 1 = \zs\" "proper_symbols zs" using 1 by auto have transforms_MI [transforms_intros]: "transforms M (snd ?cfg) (T (length y)) tps" using 2 by simp define tps' where "tps' = tps[1 := if zs = [\] then (\zs\, 2) else (\[\]\, 1)]" have "transforms M' (snd ?cfg) (T (length y) + (19 + (tps :#: Suc 0 + 2 * length zs))) tps'" unfolding M'_def by (tform tps: zs len_tps tps'_def) moreover have "T (length y) + (19 + (tps :#: Suc 0 + 2 * length zs)) \ T' (length y)" proof - have "tps :#: Suc 0 \ T (length y)" using 2 transforms_def transits_def start_config_def ex head_pos_le_time `2 \ k` by (smt (verit, ccfv_threshold) One_nat_def Suc_1 Suc_le_lessD leD linorder_le_less_linear order_less_le_trans prod.sel(2)) moreover have "length zs \ T (length y)" using zs 2 ex output_length_le_time by auto ultimately show ?thesis using T'_def 1 2 by simp qed ultimately have *: "transforms M' (snd ?cfg) (T' (length y)) tps'" using transforms_monotone by simp have "tps' ::: 1 = (if zs = [\] then tps ::: 1 else \[\]\)" using tps'_def len_tps zs(1) by simp also have "... = (if zs = [\] then \[\]\ else \[\]\)" using zs(1) by simp also have "... = (if string_to_symbols (fverify y) = [3] then \[3]\ else \[2]\)" using zs_def by simp also have "... = (if fverify y = [\] then \[\]\ else \[\]\)" by auto also have "... = (if fverify y = [\] then string_to_contents [\] else string_to_contents [\])" by auto also have "... = string_to_contents (if fverify y = [\] then [\] else [\])" by simp also have "... = string_to_contents (f' y)" using f'_def by auto finally have "tps' ::: 1 = string_to_contents (f' y)" . then show "\tps'. tps' ::: 1 = string_to_contents (f' y) \ transforms M' (snd ?cfg) (T' (length y)) tps'" using * by auto qed moreover have "\y. length (f' y) = 1" using f'_def by simp moreover have "(\x. x \ L \ (\u. length u = p (length x) \ f' \x, u\ = [\]))" using ex f'_def by auto ultimately show "L \ ?M" by blast qed qed section \$\mathcal{P}$ is a subset of $\NP$\label{s:TC-subseteq}\ text \ Let $L\in\mathcal{P}$ be a language and $M$ a Turing machine that decides $L$ in polynomial time. To show $L\in\NP$ we could use a TM that extracts the first element from the input $\langle x, u\rangle$ and runs $M$ on $x$. We do not have to construct such a TM explicitly because we have shown that the extraction of the first pair element is computable in polynomial time (lemma~@{thm [source] tm_first_computes}), and by assumption the characteristic function of $L$ is computable in polynomial time, too. The composition of these two functions is the verifier function required by the definition of $\NP$. And by lemma~@{thm [source] computes_composed_poly} the composition of polynomial-time functions is polynomial-time, too. \null \ theorem P_subseteq_NP: "\

\ \\

" proof fix L :: language assume "L \ \

" then obtain c where c: "L \ DTIME (\n. n ^ c)" using complexity_class_P_def by auto then obtain k G M T' where M: "turing_machine k G M" "computes_in_time k M (characteristic L) T'" "big_oh T' (\n. n ^ c)" using DTIME_def by auto then have 4: "big_oh_poly T'" using big_oh_poly_def by auto define f where "f = (\x. symbols_to_string (first (bindecode (string_to_symbols x))))" define T :: "nat \ nat" where "T = (\n. 9 + 4 * n)" have 1: "turing_machine 3 6 tm_first" by (simp add: tm_first_tm) have 2: "computes_in_time 3 tm_first f T" using f_def T_def tm_first_computes by simp have 3: "big_oh_poly T" proof - have "big_oh_poly (\n. 9)" using big_oh_poly_const by simp moreover have "big_oh_poly (\n. 4 * n)" using big_oh_poly_const big_oh_poly_prod big_oh_poly_poly[of 1] by simp ultimately show ?thesis using big_oh_poly_sum T_def by simp qed define fverify where "fverify = characteristic L \ f" then have *: "\T k G M. big_oh_poly T \ turing_machine k G M \ computes_in_time k M fverify T" using M 1 2 3 4 computes_composed_poly by simp then have fverify: "fverify \x, u\ = [x \ L]" for x u using f_def first_pair symbols_to_string_to_symbols fverify_def characteristic_def by simp define p :: "nat \ nat" where "p = (\_. 0)" then have "polynomial p" using const_polynomial by simp moreover have "\x. x \ L \ (\u. length u = p (length x) \ fverify \x, u\ = [\])" using fverify p_def by simp ultimately show "L \ \\

" using * complexity_class_NP_def by fast qed section \More about $\mathcal{P}$, $\NP$, and reducibility\label{s:TC-more}\ text \ We prove some low-hanging fruits about the concepts introduced in this chapter. None of the results are needed to show the Cook-Levin theorem. \ text \ A language can be reduced to itself by the identity function. Hence reducibility is a reflexive relation. \null \ lemma reducible_refl: "L \\<^sub>p L" proof - let ?M = "tm_id" let ?T = "\n. Suc (Suc n)" have "turing_machine 2 4 ?M" using tm_id_tm by simp moreover have "big_oh_poly ?T" proof - have "big_oh_poly (\n. n + 2)" using big_oh_poly_const big_oh_poly_id big_oh_poly_sum by blast then show ?thesis by simp qed moreover have "computes_in_time 2 ?M id ?T" using computes_id by simp moreover have "\x. x \ L \ id x \ L" by simp ultimately show "L \\<^sub>p L" using reducible_def by metis qed text \ Reducibility is also transitive. If $L_1 \leq_p L_2$ by a TM $M_1$ and $L_2 \leq_p L_3$ by a TM $M_2$ we merely have to run $M_2$ on the output of $M_1$ to show that $L_1 \leq_p L_3$. Again this is merely the composition of two polynomial-time computable functions. \ lemma reducible_trans: assumes "L\<^sub>1 \\<^sub>p L\<^sub>2" and "L\<^sub>2 \\<^sub>p L\<^sub>3" shows "L\<^sub>1 \\<^sub>p L\<^sub>3" proof - obtain k1 G1 M1 T1 f1 where 1: "turing_machine k1 G1 M1 \ big_oh_poly T1 \ computes_in_time k1 M1 f1 T1 \ (\x. x \ L\<^sub>1 \ f1 x \ L\<^sub>2)" using assms(1) reducible_def by metis moreover obtain k2 G2 M2 T2 f2 where 2: "turing_machine k2 G2 M2 \ big_oh_poly T2 \ computes_in_time k2 M2 f2 T2 \ (\x. x \ L\<^sub>2 \ f2 x \ L\<^sub>3)" using assms(2) reducible_def by metis ultimately obtain T k G M where "big_oh_poly T \ turing_machine k G M \ computes_in_time k M (f2 \ f1) T" using computes_composed_poly by metis moreover have "\x. x \ L\<^sub>1 \ f2 (f1 x) \ L\<^sub>3" using 1 2 by simp ultimately show ?thesis using reducible_def by fastforce qed text \ The usual way to show that a language is $\NP$-hard is to reduce another $\NP$-hard language to it. \ lemma ex_reducible_imp_NP_hard: assumes "NP_hard L'" and "L' \\<^sub>p L" shows "NP_hard L" using reducible_trans assms by auto text \ The converse is also true because reducibility is a reflexive relation. \ lemma NP_hard_iff_reducible: "NP_hard L \ (\L'. NP_hard L' \ L' \\<^sub>p L)" proof show "NP_hard L \ \L'. NP_hard L' \ L' \\<^sub>p L" using reducible_refl by auto show "\L'. NP_hard L' \ L' \\<^sub>p L \ NP_hard L" using ex_reducible_imp_NP_hard by blast qed lemma NP_complete_reducible: assumes "NP_complete L'" and "L \ \\

" and "L' \\<^sub>p L" shows "NP_complete L" using assms NP_complete_def reducible_trans by blast text \ In a sense the complexity class $\mathcal{P}$ is closed under reduction. \ lemma P_closed_reduction: assumes "L \ \

" and "L' \\<^sub>p L" shows "L' \ \

" proof - obtain c where c: "L \ DTIME (\n. n ^ c)" using assms(1) complexity_class_P_def by auto then obtain k G M T where M: "turing_machine k G M" "computes_in_time k M (characteristic L) T" "big_oh T (\n. n ^ c)" using DTIME_def by auto then have T: "big_oh_poly T" using big_oh_poly_def by auto obtain k' G' M' T' freduce where M': "turing_machine k' G' M'" "big_oh_poly T'" "computes_in_time k' M' freduce T'" "(\x. x \ L' \ freduce x \ L)" using reducible_def assms(2) by auto obtain T2 k2 G2 M2 where M2: "big_oh_poly T2" "turing_machine k2 G2 M2" "computes_in_time k2 M2 (characteristic L \ freduce) T2" using M T M' computes_composed_poly by metis obtain d where d: "big_oh T2 (\n. n ^ d)" using big_oh_poly_def M2(1) by auto have "characteristic L \ freduce = characteristic L'" using characteristic_def M'(4) by auto then have "\k G M T'. turing_machine k G M \ big_oh T' (\n. n ^ d) \ computes_in_time k M (characteristic L') T'" using M2 d by auto then have "L' \ DTIME (\n. n ^ d)" using DTIME_def by simp then show ?thesis using in_P_iff by auto qed text \ The next lemmas are items~2 and~3 of Theorem~2.8 of the textbook~\cite{ccama}. Item~1 is the transitivity of the reduction, already proved in lemma @{thm [source] reducible_trans}. \ lemma P_eq_NP: assumes "NP_hard L" and "L \ \

" shows "\

= \\

" using assms P_closed_reduction P_subseteq_NP by auto lemma NP_complete_imp: assumes "NP_complete L" shows "L \ \

\ \

= \\

" using assms NP_complete_def P_eq_NP by auto end diff --git a/thys/Cook_Levin/Oblivious.thy b/thys/Cook_Levin/Oblivious.thy --- a/thys/Cook_Levin/Oblivious.thy +++ b/thys/Cook_Levin/Oblivious.thy @@ -1,1635 +1,1633 @@ chapter \Obliviousness\label{s:oblivious}\ text \ In order to show that \SAT{} is $\NP$-hard we will eventually show how to reduce an arbitrary language $L \in \NP$ to \SAT{}. The proof can only use properties of $L$ common to all languages in $\NP$. The definition of $\NP$ provides us with a verifier Turing machine $M$ for $L$, of which we only know that it is running in polynomial time. In addition by lemma~@{text NP_output_len_1} we can assume that $M$ outputs a single bit symbol. In this chapter we are going to show that we can make additional assumptions about $M$, namely: \begin{enumerate} \item $M$ has only two tapes. \item $M$ halts on $\langle x, u\rangle$ with the output tape head on the symbol \textbf{1} iff.\ $u$ is a certificate for $x$. \item $M$ is \emph{oblivious}, which means that on any input $x$ the head positions of $M$ on all its tapes depend only on the \emph{length} of $x$, not on the symbols in $x$~\cite[Remark~1.7]{ccama}. \end{enumerate} These additional properties will somewhat simplify the reduction of $L$ to \SAT{}, more precisely the construction of the CNF formulas (see Chapter~\ref{s:Reducing}). In order to achieve this goal we will show how to simulate any polynomial-time multi-tape TM in polynomial time on a two-tape oblivious TM that halts with the output tape head on cell~1. Given a polynomial-time $k$-tape TM $M$, the basic approach is to construct a two-tape TM that encodes the $k$ tapes of $M$ on its output tape in such a way that every cell encodes $k$ symbols of $M$ and flags for $M$'s tape heads. This is the same idea as used by Dalvit and Thiemann~\cite{Multitape_To_Singletape_TM-AFP} and originally Hartmanis and Stearns~\cite{hs65} for simulating a multi-tape TM on a single-tape TM. After all our two-tape simulator can only properly use a single tape (the output/work tape). This simulator has roughly a quadratic running time overhead and so keeps the running time polynomial. However, it is not generally an oblivious TM. To make the simulator TM oblivious, we have it initially ``format'' a section on the output tape that is long enough to hold everything $M$ is going to write and whose length only depends on the input length. To simulate one step of $M$, the simulator then sweeps its output tape head all the way from the start of the tape to the end of the formatted space and back again, moving one cell per step. During these sweeps it executes one step of the simulation of $M$. Since the size of the formatted space only depends on the input length, the simulator performs the same head movements on inputs of the same length, resulting in an oblivious behavior. Moreover, it is easy to make it halt with the output tape head on cell number~1. The formatter TM is described in Section~~\ref{s:oblivious-polynomial}. The simulator TM is then constructed in Section~\ref{s:oblivious-two-tape}. Finally Section~\ref{s:oblivious-np} states the main result of this chapter. Before any of this, however, we have to define some basic concepts surrounding obliviousness. \ section \Oblivious Turing machines\label{s:oblivious-tm}\ theory Oblivious imports Memorizing begin text \ This section provides us with the tools for showing that a Turing machine is oblivious and for combining oblivious TMs into more complex oblivious TMs. So far our analysis of Turing machines involved their semantics and running time bounds. For this we mainly used the @{const transforms} predicate, which relates a start configuration and a halting configuration and an upper bound for the running time of a TM to transit from the one configuration to the other. To deal with obliviousness, we need to look more closely and inspect the sequence of tape head positions during the TM's execution, rather than only the running time. The subsections in this section roughly correspond to Sections~\ref{s:tm-basic} to~\ref{s:tm-memorizing}. In the first subsection we introduce a predicate @{term trace} analogous to @{const transforms} and show its behavior under sequential composition of TMs and loops (we will not need branches). The next subsection shows the head position sequences for those few elementary TMs from Section~\ref{s:tm-elementary} that we need for our more complex oblivious TMs later. These constructions will also heavily use the memorization-in-states technique from Section~\ref{s:tm-memorizing}, which we adapt to this chapter's needs in the final subsection. \ subsection \Traces and head positions\ text \ In order to show that a Turing machine is oblivious we need to keep track of its head positions. Consider a machine $M$ that transits from a configuration @{term cfg1} to a configuration @{term cfg2} in $t$ steps. We call the sequence of head positions on the first two tapes a \emph{trace}. If we ignore the initial head positions, the length of a trace equals $t$. Moreover we will only consider traces where $M$ either does not halt or halts in the very last step. These two properties mean, for example, that we can simply concatenate a trace of a TM that halts and trace of another TM and get the trace of the sequential execution of both TMs. Similarly, analysing while loops is simplified by these two extra assumptions. The next predicate defines what it means for a list @{term "es :: (nat \ nat) list"} to be a trace. \ definition trace :: "machine \ config \ (nat \ nat) list \ config \ bool" where "trace M cfg1 es cfg2 \ execute M cfg1 (length es) = cfg2 \ (\i (\i 0 = fst (es ! i)) \ (\i 1 = snd (es ! i))" text \ We will consider traces for machines with more than two tapes, too, but only for auxiliary constructions in combination with the memorizing-in-states technique. Therefore our definition is limited to start configurations with two tapes. A machine is \emph{oblivious} if there is a function mapping the input length to the trace that takes the machine from the start configuration with that input to a halting configuration. \ definition oblivious :: "machine \ bool" where "oblivious M \ \e. (\zs. bit_symbols zs \ (\tps. trace M (start_config 2 zs) (e (length zs)) (length M, tps)))" lemma trace_Nil: "trace M cfg [] cfg" unfolding trace_def by simp lemma traceI: assumes "execute M (q1, tps1) (length es) = (q2, tps2)" and "\i. i < length es \ fst (execute M (q1, tps1) i) < length M" and "\i. i < length es \ execute M (q1, tps1) (Suc i) <#> 0 = fst (es ! i) \ execute M (q1, tps1) (Suc i) <#> 1 = snd (es ! i)" shows "trace M (q1, tps1) es (q2, tps2)" using trace_def assms by simp lemma traceI': assumes "execute M cfg1 (length es) = cfg2" and "\i. i < length es \ fst (execute M cfg1 i) < length M" and "\i. i < length es \ execute M cfg1 (Suc i) <#> 0 = fst (es ! i) \ execute M cfg1 (Suc i) <#> 1 = snd (es ! i)" shows "trace M cfg1 es cfg2" using trace_def assms by simp lemma trace_additive: assumes "trace M (q1, tps1) es1 (q2, tps2)" and "trace M (q2, tps2) es2 (q3, tps3)" shows "trace M (q1, tps1) (es1 @ es2) (q3, tps3)" proof (rule traceI) let ?es = "es1 @ es2" show "execute M (q1, tps1) (length (es1 @ es2)) = (q3, tps3)" using trace_def assms by (simp add: execute_additive) show "fst (execute M (q1, tps1) i) < length M" if "i < length ?es" for i proof (cases "i < length es1") case True then show ?thesis using that assms(1) trace_def by simp next case False have "execute M (q1, tps1) (length es1 + (i - length es1)) = execute M (q2, tps2) (i - length es1)" using execute_additive that assms(1) trace_def by blast then have *: "execute M (q1, tps1) i = execute M (q2, tps2) (i - length es1)" using False by simp have "i - length es1 < length es2" using that False by simp then have "fst (execute M (q2, tps2) (i - length es1)) < length M" using assms(2) trace_def by simp then show ?thesis using * by simp qed show "execute M (q1, tps1) (Suc i) <#> 0 = fst (?es ! i) \ execute M (q1, tps1) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < length es1") case True then show ?thesis using that assms(1) trace_def by (simp add: nth_append) next case False have "execute M (q1, tps1) (length es1 + (Suc i - length es1)) = execute M (q2, tps2) (Suc i - length es1)" using execute_additive that assms(1) trace_def by blast then have *: "execute M (q1, tps1) (Suc i) = execute M (q2, tps2) (Suc (i - length es1))" using False by (simp add: Suc_diff_le) have "i - length es1 < length es2" using that False by simp then have "execute M (q2, tps2) (Suc (i - length es1)) <#> 0 = fst (es2 ! (i - length es1))" and "execute M (q2, tps2) (Suc (i - length es1)) <#> 1 = snd (es2 ! (i - length es1))" using assms(2) trace_def by simp_all then show ?thesis using * by (simp add: False nth_append) qed qed lemma trace_additive': assumes "trace M cfg1 es1 cfg2" and "trace M cfg2 es2 cfg3" shows "trace M cfg1 (es1 @ es2) cfg3" using trace_additive assms by (metis prod.collapse) text \ We mostly consider traces from the start state to the halting state, for which we introduce the next predicate. \ definition traces :: "machine \ tape list \ (nat \ nat) list \ tape list \ bool" where "traces M tps1 es tps2 \ trace M (0, tps1) es (length M, tps2)" text \ The relation between @{const traces} and @{const trace} is like that between @{const transforms} and @{const transits}. \ lemma tracesI [intro]: assumes "execute M (0, tps1) (length es) = (length M, tps2)" and "\i. i < length es \ fst (execute M (0, tps1) i) < length M" and "\i. i < length es \ execute M (0, tps1) (Suc i) <#> 0 = fst (es ! i) \ execute M (0, tps1) (Suc i) <#> 1 = snd (es ! i)" shows "traces M tps1 es tps2" using traces_def trace_def assms by simp lemma traces_additive: assumes "trace M (0, tps1) es1 (0, tps2)" and "traces M tps2 es2 tps3" shows "traces M tps1 (es1 @ es2) tps3" using assms traces_def trace_additive by simp lemma execute_trace_append: assumes "trace M1 (0, tps1) es1 (length M1, tps2)" (is "trace _ ?cfg1 _ _") and "t \ length es1" shows "execute (M1 @ M2) (0, tps1) t = execute M1 (0, tps1) t" (is "execute ?M _ _ = _") using assms(2) proof (induction t) case 0 then show ?case by simp next case (Suc t) then have "t < length es1" by simp then have 1: "fst (execute M1 ?cfg1 t) < length M1" using traces_def trace_def assms(1) by simp have 2: "length ?M = length M1 + length M2" using length_turing_machine_sequential by simp have "execute ?M ?cfg1 (Suc t) = exe ?M (execute ?M ?cfg1 t)" by simp also have "... = exe ?M (execute M1 ?cfg1 t)" (is "_ = exe _ ?cfg") using Suc by simp also have "... = sem (?M ! (fst ?cfg)) ?cfg" using 1 2 exe_def by simp also have "... = sem (M1 ! (fst ?cfg)) ?cfg" using 1 by (simp add: nth_append turing_machine_sequential_def) also have "... = exe M1 (execute M1 ?cfg1 t)" using exe_def 1 by simp also have "... = execute M1 ?cfg1 (Suc t)" by simp finally show ?case . qed subsection \Increasing the number of tapes\ text \ This is lemma @{thm [source] transforms_append_tapes} adapted for @{const traces}. \ lemma traces_append_tapes: assumes "turing_machine 2 G M" and "length tps1 = 2" and "traces M tps1 es tps2" shows "traces (append_tapes 2 (2 + length tps') M) (tps1 @ tps') es (tps2 @ tps')" proof let ?M = "append_tapes 2 (2 + length tps') M" show "execute ?M (0, tps1 @ tps') (length es) = (length ?M, tps2 @ tps')" proof - have "execute M (0, tps1) (length es) = (length M, tps2)" using assms(3) by (simp add: trace_def traces_def) moreover have "execute ?M (0, tps1 @ tps') (length es) = (fst (execute M (0, tps1) (length es)), snd (execute M (0, tps1) (length es)) @ tps')" using execute_append_tapes'[OF assms(1-2)] by simp ultimately show ?thesis by (simp add: length_append_tapes) qed show "fst (execute ?M (0, tps1 @ tps') i) < length ?M" if "i < length es" for i proof - have "fst (execute M (0, tps1) i) < length M" using that assms(3) trace_def traces_def by blast then show "fst (execute ?M (0, tps1 @ tps') i) < length ?M" by (metis (no_types) assms(1,2) execute_append_tapes' fst_conv length_append_tapes) qed show "snd (execute ?M (0, tps1 @ tps') (Suc i)) :#: 0 = fst (es ! i) \ snd (execute ?M (0, tps1 @ tps') (Suc i)) :#: 1 = snd (es ! i)" if "i < length es" for i proof - have "snd (execute ?M (0, tps1 @ tps') (Suc i)) = snd (execute M (0, tps1) (Suc i)) @ tps'" using execute_append_tapes' assms by (metis snd_conv) moreover have "||execute M (0, tps1) (Suc i)|| = 2" using assms(1,2) by (metis execute_num_tapes snd_conv) ultimately show ?thesis using that assms by (simp add: nth_append trace_def traces_def) qed qed subsection \Combining Turing machines\ text \ Traces for sequentially composed Turing machines are just concatenated traces of the individual machines. \ lemma traces_sequential: assumes "traces M1 tps1 es1 tps2" and "traces M2 tps2 es2 tps3" shows "traces (M1 ;; M2) tps1 (es1 @ es2) tps3" proof let ?M = "M1 ;; M2" let ?cfg1 = "(0, tps1)" let ?cfg1' = "(length M1, tps2)" let ?cfg2 = "(0, tps2)" let ?cfg2' = "(length M2, tps3)" let ?es = "es1 @ es2" have 3: "execute M1 ?cfg1 (length es1) = ?cfg1'" using assms(1) traces_def trace_def by simp have "fst ?cfg1 = 0" by simp have 4: "execute M2 ?cfg2 (length es2) = ?cfg2'" using assms(2) traces_def trace_def by auto have "?cfg1' = ?cfg2 <+=> length M1" by simp have 2: "length ?M = length M1 + length M2" using length_turing_machine_sequential by simp have t_le: "execute ?M ?cfg1 t = execute M1 ?cfg1 t" if "t \ length es1" for t using that proof (induction t) case 0 then show ?case by simp next case (Suc t) then have "t < length es1" by simp then have 1: "fst (execute M1 ?cfg1 t) < length M1" using traces_def trace_def assms(1) by simp have "execute ?M ?cfg1 (Suc t) = exe ?M (execute ?M ?cfg1 t)" by simp also have "... = exe ?M (execute M1 ?cfg1 t)" (is "_ = exe _ ?cfg") using Suc by simp also have "... = sem (?M ! (fst ?cfg)) ?cfg" using 1 2 exe_def by simp also have "... = sem (M1 ! (fst ?cfg)) ?cfg" using 1 by (simp add: nth_append turing_machine_sequential_def) also have "... = exe M1 (execute M1 ?cfg1 t)" using exe_def 1 by simp also have "... = execute M1 ?cfg1 (Suc t)" by simp finally show ?case . qed have t_ge: "execute ?M ?cfg1 (length es1 + t) = execute M2 ?cfg2 t <+=> length M1" if "t \ length es2" for t using that proof (induction t) case 0 then show ?case using t_le 3 by simp next case (Suc t) have "execute ?M ?cfg1 (length es1 + Suc t) = execute ?M ?cfg1 (Suc (length es1 + t))" by simp also have "... = exe ?M (execute ?M ?cfg1 (length es1 + t))" by simp also have "... = exe ?M (execute M2 ?cfg2 t <+=> length M1)" (is "_ = exe _ (?cfg <+=> _)") using Suc by simp also have "... = (exe M2 (execute M2 ?cfg2 t)) <+=> length M1" using exe_relocate by simp also have "... = execute M2 ?cfg2 (Suc t) <+=> length M1" by simp finally show ?case . qed show "fst (execute ?M ?cfg1 i) < length ?M" if "i < length ?es" for i proof (cases "i < length es1") case True then show ?thesis using t_le assms(1) traces_def trace_def 2 by auto next case False then obtain i' where "i = length es1 + i'" "i' \ length es2" by (metis \i < length (es1 @ es2)\ add_diff_inverse_nat add_le_cancel_left length_append less_or_eq_imp_le) then show ?thesis using t_ge assms(2) traces_def trace_def that 2 by simp qed show "execute ?M ?cfg1 (length ?es) = (length ?M, tps3)" by (simp add: 2 4 t_ge) show "execute ?M ?cfg1 (Suc i) <#> 0 = fst (?es ! i) \ execute ?M ?cfg1 (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < length es1") case True then have "Suc i \ length es1" by simp then have "execute ?M ?cfg1 (Suc i) = execute M1 ?cfg1 (Suc i)" using t_le by blast then show ?thesis using assms(1) traces_def trace_def by (simp add: True nth_append) next case False have 8: "i - length es1 < length es2" using False that by simp with False have "Suc i - length es1 \ length es2" by simp then have "execute ?M ?cfg1 (Suc i) = execute M2 ?cfg2 (Suc i - length es1) <+=> length M1" using t_ge False by fastforce moreover have "?es ! i = es2 ! (i - length es1)" by (simp add: False nth_append) moreover have "execute M2 ?cfg2 (Suc i) <#> 0 = fst (es2 ! i) \ execute M2 ?cfg2 (Suc i) <#> 1 = snd (es2 ! i)" if "i < length es2" for i using that assms(2) traces_def trace_def by simp ultimately show ?thesis by (metis "8" False Nat.add_diff_assoc le_less_linear plus_1_eq_Suc snd_conv) qed qed text \ Next we show how to derive traces for machines created by the @{text WHILE} operation. If the condition is false, the trace of the loop is the trace for the machine computing the condition plus a singleton trace for the jump. \ lemma tm_loop_sem_false_trace: assumes "traces M1 tps0 es1 tps1" and "\ cond (read tps1)" shows "trace (WHILE M1 ; cond DO M2 DONE) (0, tps0) (es1 @ [(tps1 :#: 0, tps1 :#: 1)]) (length M1 + length M2 + 2, tps1)" (is "trace ?M _ _ _") proof (rule traceI) let ?C1 = "M1" let ?C2 = "[cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)]" let ?C3 = "relocate (length M1 + 1) M2" let ?C4 = "[cmd_jump (\_. True) 0 0]" let ?C34 = "?C3 @ ?C4" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_loop_def by simp then have 1: "?M ! (length M1) = cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)" by simp let ?es = "es1 @ [(tps1 :#: 0, tps1 :#: 1)]" show goal1: "execute ?M (0, tps0) (length ?es) = (length M1 + length M2 + 2, tps1)" proof - have "execute ?M (0, tps0) (length es1) = execute M1 (0, tps0) (length es1)" using execute_trace_append assms by (simp add: traces_def turing_machine_loop_def) then have 2: "execute ?M (0, tps0) (length es1) = (length M1, tps1) " using assms trace_def traces_def by simp have "execute ?M (0, tps0) (length ?es) = execute ?M (0, tps0) (Suc (length es1))" by simp also have "... = exe ?M (execute ?M (0, tps0) (length es1))" by simp also have "... = exe ?M (length M1, tps1)" using 2 by simp also have "... = sem (cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)) (length M1, tps1)" by (simp add: "1" exe_lt_length turing_machine_loop_len) also have "... = (length M1 + length M2 + 2, tps1)" using assms(2) sem_jump by simp finally show ?thesis . qed show "fst (execute ?M (0, tps0) i) < length ?M" if "i < length ?es" for i proof (cases "i < length es1") case True then have "execute ?M (0, tps0) i = execute M1 (0, tps0) i" using execute_trace_append assms parts by (simp add: traces_def) then show ?thesis using assms(1) trace_def traces_def True turing_machine_loop_len by auto next case False with that have "i = length es1" by simp then show ?thesis using assms(1) trace_def traces_def turing_machine_loop_len by (simp add: execute_trace_append parts) qed show "execute ?M (0, tps0) (Suc i) <#> 0 = fst (?es ! i) \ execute ?M (0, tps0) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < length es1") case True then have "Suc i \ length es1" by simp then have "execute ?M (0, tps0) (Suc i) = execute M1 (0, tps0) (Suc i)" using execute_trace_append assms parts by (metis traces_def) then show ?thesis using assms(1) trace_def traces_def True by (simp add: nth_append) next case False with that have "Suc i = length ?es" by simp then show ?thesis using goal1 by simp qed qed lemma tm_loop_sem_false_traces: assumes "traces M1 tps0 es1 tps1" and "\ cond (read tps1)" and "es = es1 @ [(tps1 :#: 0, tps1 :#: 1)]" shows "traces (WHILE M1 ; cond DO M2 DONE) tps0 es tps1" using tm_loop_sem_false_trace assms traces_def turing_machine_loop_len by fastforce text \ If the loop condition evaluates to true, the trace of one iteration is the concatenation of the traces of the condition machine and the loop body machine with two additional singleton traces for the jumps. \ lemma tm_loop_sem_true_traces: assumes "traces M1 tps0 es1 tps1" and "traces M2 tps1 es2 tps2" and "cond (read tps1)" shows "trace (WHILE M1 ; cond DO M2 DONE) (0, tps0) (es1 @ [(tps1 :#: 0, tps1 :#: 1)] @ es2 @ [(tps2 :#: 0, tps2 :#: 1)]) (0, tps2)" (is "trace ?M _ ?es _") proof (rule traceI) let ?C1 = "M1" let ?C2 = "[cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)]" let ?C3 = "relocate (length M1 + 1) M2" let ?C4 = "[cmd_jump (\_. True) 0 0]" let ?C34 = "?C3 @ ?C4" have parts: "?M = ?C1 @ ?C2 @ ?C3 @ ?C4" using turing_machine_loop_def by simp then have 1: "?M ! (length M1) = cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)" by simp from parts have parts': "?M = ((?C1 @ ?C2) @ ?C3) @ ?C4" by simp have len_M: "length ?M = length M1 + length M2 + 2" using turing_machine_loop_len assms by simp have len_es: "length ?es = length es1 + length es2 + 2" by simp have exec_1: "execute ?M (0, tps0) t = execute M1 (0, tps0) t" if "t \ length es1" for t using execute_trace_append assms by (simp add: parts that traces_def) have exec_2: "execute ?M (0, tps0) (length es1 + 1) = (length M1 + 1, tps1)" proof - have "execute ?M (0, tps0) (length es1) = execute M1 (0, tps0) (length es1)" using execute_trace_append assms by (simp add: traces_def turing_machine_loop_def) then have 2: "execute ?M (0, tps0) (length es1) = (length M1, tps1) " using assms trace_def traces_def by simp have "execute ?M (0, tps0) (length es1 + 1) = execute ?M (0, tps0) (Suc (length es1))" by simp also have "... = exe ?M (execute ?M (0, tps0) (length es1))" by simp also have "... = exe ?M (length M1, tps1)" using 2 by simp also have "... = sem (cmd_jump cond (length M1 + 1) (length M1 + length M2 + 2)) (length M1, tps1)" by (simp add: "1" exe_lt_length turing_machine_loop_len) also have "... = (length M1 + 1, tps1)" using assms(3) sem_jump by simp finally show ?thesis . qed have exec_3': "execute ?M (0, tps0) (length es1 + 1 + t) = execute M2 (0, tps1) t <+=> (length M1 + 1)" if "t \ length es2" for t using that proof (induction t) case 0 then show ?case using exec_2 by simp next case (Suc t) then have 2: "fst (execute M2 (0, tps1) t) < length M2" using assms(2) trace_def traces_def by simp then have 3: "fst (execute M2 (0, tps1) t <+=> (length M1 + 1)) < length M1 + length M2 + 1" by simp have 4: "fst (execute M2 (0, tps1) t <+=> (length M1 + 1)) \ length M1 + 1" by simp have "?M = (?C1 @ ?C2) @ (?C3 @ ?C4)" using parts by simp moreover have "length (?C1 @ ?C2) = length M1 + 1" by simp ultimately have "?M ! i = (?C3 @ ?C4) ! (i - (length M1 + 1))" if "i \ length M1 + 1" and "i < length M1 + length M2 + 1" for i using that by (simp add: nth_append) then have "?M ! i = ?C3 ! (i - (length M1 + 1))" if "i \ length M1 + 1" and "i < length M1 + length M2 + 1" for i - using that - by (smt One_nat_def \length (?C1 @ ?C2) = length M1 + 1\ - add.right_neutral add_Suc_right append.assoc length_append not_le nth_append plus_nat.simps(2) length_relocate) + using that by (simp add: length_relocate less_diff_conv2 nth_append) with 3 4 have "?M ! (fst (execute M2 (0, tps1) t <+=> (length M1 + 1))) = ?C3 ! ((fst (execute M2 (0, tps1) t <+=> (length M1 + 1))) - (length M1 + 1))" by simp then have in_C3: "?M ! (fst (execute M2 (0, tps1) t <+=> (length M1 + 1))) = ?C3 ! ((fst (execute M2 (0, tps1) t)))" by simp have "execute ?M (0, tps0) (length es1 + 1 + Suc t) = execute ?M (0, tps0) (Suc (length es1 + 1 + t))" by simp also have "... = exe ?M (execute ?M (0, tps0) (length es1 + 1 + t))" by simp also have "... = exe ?M (execute M2 (0, tps1) t <+=> (length M1 + 1))" (is "_ = exe ?M ?cfg") using Suc by simp also have "... = sem (?M ! (fst ?cfg)) ?cfg" using exe_def "3" len_M by simp also have "... = sem (?C3 ! (fst (execute M2 (0, tps1) t))) (execute M2 (0, tps1) t)" using in_C3 sem by simp also have "... = sem (M2 ! (fst (execute M2 (0, tps1) t))) (execute M2 (0, tps1) t) <+=> (length M1 + 1)" using sem_relocate 2 by simp also have "... = exe M2 (execute M2 (0, tps1) t) <+=> (length M1 + 1)" by (simp add: 2 exe_def) also have "... = (execute M2 (0, tps1) (Suc t)) <+=> (length M1 + 1)" by simp finally show ?case . qed then have exec_3: "execute ?M (0, tps0) t = execute M2 (0, tps1) (t - (length es1 + 1)) <+=> (length M1 + 1)" if "t \ length es1 + 1 " and "t \ length es1 + length es2 + 1" for t using that - by (smt Nat.add_diff_assoc2 Nat.diff_diff_right add_diff_cancel_left' add_diff_cancel_right' le_Suc_ex le_add2) + by (smt (verit) Nat.add_diff_assoc2 Nat.diff_diff_right add_diff_cancel_left' add_diff_cancel_right' le_Suc_ex le_add2) have exec_4: "execute ?M (0, tps0) (length es1 + length es2 + 2) = (0, tps2)" proof - have "execute ?M (0, tps0) (length es1 + length es2 + 2) = execute ?M (0, tps0) (Suc (length es1 + length es2 + 1))" by simp also have "... = exe ?M (execute ?M (0, tps0) (length es1 + length es2 + 1))" by simp also have "... = exe ?M (execute M2 (0, tps1) (length es2) <+=> (length M1 + 1))" (is "_ = exe ?M ?cfg") using exec_3' by simp also have "... = sem (?M ! (fst ?cfg)) ?cfg" using exe_def assms(2) len_M trace_def traces_def by auto also have "... = sem (cmd_jump (\_. True) 0 0) ?cfg" proof - have "fst ?cfg = length M1 + length M2 + 1" using assms(2) len_M trace_def traces_def by simp then have "?M ! (fst ?cfg) = cmd_jump (\_. True) 0 0" by (metis (no_types, lifting) add.right_neutral add_Suc_right length_Cons list.size(3) nth_append_length nth_append_length_plus parts plus_1_eq_Suc length_relocate) then show ?thesis by simp qed also have "... = (0, tps2)" using assms(2) sem_jump trace_def traces_def by auto finally show ?thesis by simp qed show "execute ?M (0, tps0) (length ?es) = (0, tps2)" using exec_4 by auto show " fst (execute ?M (0, tps0) i) < length ?M" if "i < length ?es" for i proof - consider "i < length es1" | "i = length es1" | "i \ length es1 + 1 " and "i \ length es1 + length es2 + 1" | "i = length es1 + length es2 + 2" using `i < length ?es` by fastforce then show ?thesis proof (cases) case 1 then have "fst (execute ?M (0, tps0) i) = fst (execute M1 (0, tps0) i)" using exec_1 by simp moreover have "\i (length M1 + 1)" using exec_3 by simp have a: "\i length es2" using 3 by simp then have "fst (execute M2 (0, tps1) (i - (length es1 + 1))) \ length M2" using a b that using le_eq_less_or_eq by auto then have "fst (execute M2 (0, tps1) (i - (length es1 + 1)) <+=> (length M1 + 1)) < length ?M" by (simp add: len_M) then show ?thesis using eq by simp next case 4 then show ?thesis using exec_4 using len_es that by linarith qed qed show "execute ?M (0, tps0) (Suc i) <#> 0 = fst (?es ! i) \ execute ?M (0, tps0) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof - consider "i < length es1" | "i = length es1" | "i \ length es1 + 1 " and "i < length es1 + length es2 + 1" | "i = length es1 + length es2 + 1" using `i < length ?es` by fastforce then show ?thesis proof (cases) case 1 then have "Suc i \ length es1" by simp then have "execute ?M (0, tps0) (Suc i) = execute M1 (0, tps0) (Suc i)" using exec_1 by blast then show ?thesis using assms(1) trace_def traces_def by (simp add: "1" nth_append) next case 2 then have "execute ?M (0, tps0) (Suc i) = (length M1 + 1, tps1)" using exec_2 by simp then show ?thesis using 2 by simp next case 3 then have Suc_i: "Suc i \ length es1 + 1 " "Suc i \ length es1 + length es2 + 1" by simp_all then have *: "execute ?M (0, tps0) (Suc i) = execute M2 (0, tps1) (Suc i - (length es1 + 1)) <+=> (length M1 + 1)" using exec_3 by blast from 3 have i: "i - (length es1 + 1) < length es2" (is "?j < length es2") by simp then have **: "execute M2 (0, tps1) (Suc ?j) <#> 0 = fst (es2 ! ?j) \ execute M2 (0, tps1) (Suc ?j) <#> 1 = snd (es2 ! ?j)" using assms(2) trace_def traces_def by simp have "((es1 @ [(tps1 :#: 0, tps1 :#: 1)]) @ es2) ! i = es2 ! ?j" using i 3 by (simp add: nth_append) then have "es2 ! ?j = ?es ! i" by (metis Suc_eq_plus1 append.assoc i length_append_singleton nth_append) then show ?thesis using * ** using "3"(1) Suc_diff_le by fastforce next case 4 then have "execute ?M (0, tps0) (Suc i) = (0, tps2)" using exec_4 by simp then show ?thesis by (simp add: "4" nth_append) qed qed qed lemma tm_loop_sem_true_tracesI: assumes "traces M1 tps0 es1 tps1" and "traces M2 tps1 es2 tps2" and "cond (read tps1)" and "es = es1 @ [(tps1 :#: 0, tps1 :#: 1)] @ es2 @ [(tps2 :#: 0, tps2 :#: 1)]" shows "trace (WHILE M1 ; cond DO M2 DONE) (0, tps0) es (0, tps2)" using assms tm_loop_sem_true_traces by blast text \ Combining traces for $m$ iterations of a loop. Typically $m$ will be the total number of iterations. \ lemma tm_loop_trace_simple: fixes m :: nat and M :: machine and tps :: "nat \ tape list" and es :: "nat \ (nat \ nat) list" assumes "\i. i < m \ trace M (0, tps i) (es i) (0, tps (Suc i))" shows "trace M (0, tps 0) (concat (map es [0.. For simple loops, where we have an upper bound for the length of traces independent of the iteration, there is a trivial upper bound for the length of the trace of $m$ iterations. This is the only situation we will encounter. \ lemma length_concat_le: assumes "\i. i < m \ length (es i) \ b" shows "length (concat (map es [0.. m * b" using assms proof (induction m) case 0 then show ?case by simp next case (Suc m) have "length (concat (map es [0.. m * b + length (es m)" using Suc by simp also have "... \ m * b + b" using Suc by simp also have "... = (Suc m) * b" by simp finally show ?case . qed subsection \Traces for elementary Turing machines\label{s:oblivious-traces}\ text \ Just like the not necessarily oblivious Turing machines considered so far, our oblivious Turing machines will be built from elementary ones from Section~\ref{s:tm-elementary}. In this subsection we show the traces of all the elementary machines we will need. \ lemma tm_left_0_traces: assumes "length tps > 1" shows "traces (tm_left 0) tps [(tps :#: 0 - 1, tps :#: 1)] (tps[0:=(fst (tps ! 0), snd (tps ! 0) - 1)])" proof - from assms have "length tps > 0" by auto with assms show ?thesis using execute_tm_left assms tm_left_def by (intro tracesI) simp_all qed lemma traces_tm_left_0I: assumes "length tps > 1" and "es = [(tps :#: 0 - 1, tps :#: 1)]" and "tps' = (tps[0:=(fst (tps ! 0), snd (tps ! 0) - 1)])" shows "traces (tm_left 0) tps es tps'" using tm_left_0_traces assms by simp lemma tm_left_1_traces: assumes "length tps > 1" shows "traces (tm_left 1) tps [(tps :#: 0, tps :#: 1 - 1)] (tps[1:=(fst (tps ! 1), snd (tps ! 1) - 1)])" proof - from assms have "length tps > 0" by auto with assms show ?thesis using execute_tm_left assms tm_left_def by (intro tracesI) simp_all qed lemma traces_tm_left_1I: assumes "length tps > 1" and "es = [(tps :#: 0, tps :#: 1 - 1)]" and "tps' = (tps[1:=(fst (tps ! 1), snd (tps ! 1) - 1)])" shows "traces (tm_left 1) tps es tps'" using tm_left_1_traces assms by simp lemma tm_right_0_traces: assumes "length tps > 1" shows "traces (tm_right 0) tps [(tps :#: 0 + 1, tps :#: 1)] (tps[0:=(fst (tps ! 0), snd (tps ! 0) + 1)])" proof - from assms have "length tps > 0" by auto with assms show ?thesis using execute_tm_right assms tm_right_def by (intro tracesI) simp_all qed lemma traces_tm_right_0I: assumes "length tps > 1" and "es = [(tps :#: 0 + 1, tps :#: 1)]" and "tps' = (tps[0:=(fst (tps ! 0), snd (tps ! 0) + 1)])" shows "traces (tm_right 0) tps es tps'" using tm_right_0_traces assms by simp lemma tm_right_1_traces: assumes "length tps > 1" shows "traces (tm_right 1) tps [(tps :#: 0, tps :#: 1 + 1)] (tps[1:=(fst (tps ! 1), snd (tps ! 1) + 1)])" proof - from assms have "length tps > 0" by auto with assms show ?thesis using execute_tm_right assms tm_right_def by (intro tracesI) simp_all qed lemma tm_rtrans_1_traces: assumes "1 < length tps" shows "traces (tm_rtrans 1 f) tps [(tps :#: 0, tps :#: 1 + 1)] (tps[1 := tps ! 1 |:=| f (tps :.: 1) |+| 1])" using execute_tm_rtrans assms tm_rtrans_def by (intro tracesI) simp_all lemma traces_tm_right_1I: assumes "length tps > 1" and "es = [(tps :#: 0, tps :#: 1 + 1)]" and "tps' = (tps[1:=(fst (tps ! 1), snd (tps ! 1) + 1)])" shows "traces (tm_right 1) tps es tps'" using tm_right_1_traces assms by simp lemma traces_tm_rtrans_1I: assumes "1 < length tps" and "es = [(tps :#: 0, tps :#: 1 + 1)]" and "tps' = (tps[1 := tps ! 1 |:=| f (tps :.: 1) |+| 1])" shows "traces (tm_rtrans 1 f) tps es tps'" using tm_rtrans_1_traces assms by simp lemma tm_left_until_1_traces: assumes "length tps > 1" and "begin_tape H (tps ! 1)" shows "traces (tm_left_until H 1) tps (map (\i. (tps :#: 0, i)) (rev [0..i. i < length ?es \ fst (execute (tm_left_until H 1) (0, tps) i) < length (tm_left_until H 1)" using execute_tm_left_until_less assms tm_left_until_def by simp show "execute (tm_left_until H 1) (0, tps) (Suc i) <#> 0 = fst (?es ! i) \ execute (tm_left_until H 1) (0, tps) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < tps :#: 1") case True then have i: "Suc i \ tps :#: 1" by simp then have "execute (tm_left_until H 1) (0, tps) (Suc i) = (0, tps[1 := tps ! 1 |-| Suc i])" using execute_tm_left_until_less assms by presburger moreover have "?es ! i = (tps :#: 0, tps :#: 1 - Suc i)" proof - have "?es ! i = (map (\i. (tps :#: 0, i)) (rev [0.. 1" and "begin_tape H (tps ! 1)" and "es = map (\i. (tps :#: 0, i)) (rev [0.. 1" and "begin_tape H (tps ! 0)" shows "traces (tm_left_until H 0) tps (map (\i. (i, tps :#: 1)) (rev [0.. 0" using assms(1) by auto let ?es = "map (\i. (i, tps :#: 1)) (rev [0..i. i < length ?es \ fst (execute (tm_left_until H 0) (0, tps) i) < length (tm_left_until H 0)" using execute_tm_left_until_less len assms(2) tm_left_until_def by simp show "execute (tm_left_until H 0) (0, tps) (Suc i) <#> 0 = fst (?es ! i) \ execute (tm_left_until H 0) (0, tps) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < tps :#: 0") case True then have i: "Suc i \ tps :#: 0" by simp then have "execute (tm_left_until H 0) (0, tps) (Suc i) = (0, tps[0 := tps ! 0 |-| Suc i])" using execute_tm_left_until_less assms(2) len by blast moreover have "?es ! i = (tps :#: 0 - Suc i, tps :#: 1)" proof - have "?es ! i = (map (\i. (i, tps :#: 1)) (rev [0.. 1" and "begin_tape H (tps ! 0)" and "es = map (\i. (i, tps :#: 1)) (rev [0.. 1" and "clean_tape (tps ! 0)" shows "traces (tm_start 0) tps (map (\i. (i, tps :#: 1)) (rev [0.. 1" and "clean_tape (tps ! 1)" shows "traces (tm_start 1) tps (map (\i. (tps :#: 0, i)) (rev [0.. 1" and "clean_tape (tps ! 1)" and "es = map (\i. (tps :#: 0, i)) (rev [0.. 1" and "clean_tape (tps ! 0)" shows "traces (tm_cr 0) tps ((map (\i. (i, tps :#: 1)) (rev [0.. 0" by auto show "traces (tm_start 0) tps (map (\i. (i, tps :#: 1)) (rev [0.. 1" and "clean_tape (tps ! 0)" and "es = map (\i. (i, tps :#: 1)) (rev [0.. 1" and "clean_tape (tps ! 1)" shows "traces (tm_cr 1) tps ((map (\i. (tps :#: 0, i)) (rev [0.. 1" and "clean_tape (tps ! 1)" and "es = map (\i. (tps :#: 0, i)) (rev [0.. n" shows "execute (tm_trans_until j1 j2 H f) (0, tps) t <#> j1 = tps :#: j1 + t" and "execute (tm_trans_until j1 j2 H f) (0, tps) t <#> j2 = tps :#: j2 + t" using assms execute_tm_trans_until_less[OF assms] by ((metis (no_types, lifting) length_list_update nth_list_update_eq nth_list_update_neq sndI transplant_def), (metis (no_types, lifting) length_list_update nth_list_update_eq sndI transplant_def)) lemma heads_tm_ltrans_until_less: assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "t \ n" and "n \ tps :#: j1" and "n \ tps :#: j2" shows "execute (tm_ltrans_until j1 j2 H f) (0, tps) t <#> j1 = tps :#: j1 - t" and "execute (tm_ltrans_until j1 j2 H f) (0, tps) t <#> j2 = tps :#: j2 - t" using assms execute_tm_ltrans_until_less[OF assms] by ((metis (no_types, lifting) length_list_update nth_list_update_eq nth_list_update_neq sndI ltransplant_def), (metis (no_types, lifting) length_list_update nth_list_update_eq sndI ltransplant_def)) lemma heads_tm_trans_until_less': assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" and "t \ n" and "j \ j1" and "j \ j2" shows "execute (tm_trans_until j1 j2 H f) (0, tps) t <#> j = tps :#: j" using assms execute_tm_trans_until_less[OF assms(1-5)] by simp lemma heads_tm_ltrans_until_less': assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "t \ n" and "n \ tps :#: j1" and "n \ tps :#: j2" and "j \ j1" and "j \ j2" shows "execute (tm_ltrans_until j1 j2 H f) (0, tps) t <#> j = tps :#: j" using assms execute_tm_ltrans_until_less[OF assms(1-7)] by simp lemma heads_tm_trans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" shows "execute (tm_trans_until j1 j2 H f) (0, tps) (Suc n) <#> j1 = tps :#: j1 + n" and "execute (tm_trans_until j1 j2 H f) (0, tps) (Suc n) <#> j2 = tps :#: j2 + n" using assms execute_tm_trans_until[OF assms] by ((metis (no_types, lifting) length_list_update nth_list_update_eq nth_list_update_neq snd_conv transplant_def), (metis (no_types, lifting) length_list_update nth_list_update_eq snd_conv transplant_def)) lemma heads_tm_ltrans_until: assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "n \ tps :#: j1" and "n \ tps :#: j2" shows "execute (tm_ltrans_until j1 j2 H f) (0, tps) (Suc n) <#> j1 = tps :#: j1 - n" and "execute (tm_ltrans_until j1 j2 H f) (0, tps) (Suc n) <#> j2 = tps :#: j2 - n" using assms execute_tm_ltrans_until[OF assms] by ((metis (no_types, lifting) length_list_update nth_list_update_eq nth_list_update_neq snd_conv ltransplant_def), (metis (no_types, lifting) length_list_update nth_list_update_eq snd_conv ltransplant_def)) lemma heads_tm_trans_until': assumes "j1 < k" and "j2 < k" and "length tps = k" and "rneigh (tps ! j1) H n" and "j \ j1" and "j \ j2" shows "execute (tm_trans_until j1 j2 H f) (0, tps) (Suc n) <#> j = tps :#: j" using assms execute_tm_trans_until[OF assms(1-4)] by simp lemma heads_tm_ltrans_until': assumes "j1 < k" and "j2 < k" and "length tps = k" and "lneigh (tps ! j1) H n" and "n \ tps :#: j1" and "n \ tps :#: j2" and "j \ j1" and "j \ j2" shows "execute (tm_ltrans_until j1 j2 H f) (0, tps) (Suc n) <#> j = tps :#: j" using assms execute_tm_ltrans_until[OF assms(1-6)] by simp lemma traces_tm_trans_until_11: assumes "1 < k" and "length tps = k" and "rneigh (tps ! 1) H n" shows "traces (tm_trans_until 1 1 H f) tps (map (\i. (tps :#: 0, tps :#: 1 + Suc i)) [0.. n" by simp then have "fst (execute ?M (0, tps) i) = 0" using execute_tm_trans_until_less assms by simp then show ?thesis using tm_trans_until_def by simp qed show "execute ?M (0, tps) (Suc i) <#> 0 = fst (?es ! i) \ execute ?M (0, tps) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < n") case True then have "?es ! i = (tps :#: 0, tps :#: 1 + Suc i)" by (simp add: nth_append) moreover from True have "Suc i \ n" by simp ultimately show ?thesis using heads_tm_trans_until_less' heads_tm_trans_until_less assms by (metis One_nat_def Suc_neq_Zero fst_conv snd_conv) next case False then have "i = n" using that by simp then have "?es ! i = (tps :#: 0, tps :#: 1 + n)" by (metis (no_types, lifting) diff_zero length_map length_upt nth_append_length) then show ?thesis using heads_tm_trans_until' heads_tm_trans_until assms `i = n` by simp qed qed lemma traces_tm_ltrans_until_11: assumes "1 < k" and "length tps = k" and "lneigh (tps ! 1) H n" and "n \ tps :#: 1" shows "traces (tm_ltrans_until 1 1 H f) tps (map (\i. (tps :#: 0, tps :#: 1 - Suc i)) [0.. n" by simp then have "fst (execute ?M (0, tps) i) = 0" using execute_tm_ltrans_until_less assms by simp then show ?thesis using tm_ltrans_until_def by simp qed show "execute ?M (0, tps) (Suc i) <#> 0 = fst (?es ! i) \ execute ?M (0, tps) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < n") case True then have "?es ! i = (tps :#: 0, tps :#: 1 - Suc i)" by (simp add: nth_append) moreover from True have "Suc i \ n" by simp ultimately show ?thesis using heads_tm_ltrans_until_less' heads_tm_ltrans_until_less assms by (metis One_nat_def Suc_neq_Zero fst_conv snd_conv) next case False then have "i = n" using that by simp then have "?es ! i = (tps :#: 0, tps :#: 1 - n)" by (metis (no_types, lifting) diff_zero length_map length_upt nth_append_length) then show ?thesis using heads_tm_ltrans_until' heads_tm_ltrans_until assms `i = n` by simp qed qed lemma traces_tm_trans_until_01: assumes "0 < k" and "1 < k" and "length tps = k" and "rneigh (tps ! 0) H n" shows "traces (tm_trans_until 0 1 H f) tps (map (\i. (tps :#: 0 + Suc i, tps :#: 1 + Suc i)) [0.. n" by simp then have "fst (execute ?M (0, tps) i) = 0" using execute_tm_trans_until_less[of 0 k 1] assms by simp then show ?thesis using tm_trans_until_def by simp qed show "execute ?M (0, tps) (Suc i) <#> 0 = fst (?es ! i) \ execute ?M (0, tps) (Suc i) <#> 1 = snd (?es ! i)" if "i < length ?es" for i proof (cases "i < n") case True then have "?es ! i = (tps :#: 0 + Suc i, tps :#: 1 + Suc i)" by (simp add: nth_append) moreover from True have "Suc i \ n" by simp ultimately show ?thesis using heads_tm_trans_until_less[of 0 k 1 tps H n "Suc i" f] assms by simp next case False then have "i = n" using that by simp then have "?es ! i = (tps :#: 0 + n, tps :#: 1 + n)" by (metis (no_types, lifting) diff_zero length_map length_upt nth_append_length) then show ?thesis using heads_tm_trans_until[of 0 k 1 tps H n f] assms `i = n` by simp qed qed lemma traces_tm_trans_until_01I: assumes "1 < length tps" and "rneigh (tps ! 0) H n" and "es = map (\i. (tps :#: 0 + Suc i, tps :#: 1 + Suc i)) [0..i. (tps :#: 0, tps :#: 1 + Suc i)) [0..h tps :#: 1" and "es = map (\i. (tps :#: 0, tps :#: 1 - Suc i)) [0..i. (tps :#: 0 + Suc i, tps :#: 1 + Suc i)) [0..i. (tps :#: 0, tps :#: 1 + Suc i)) [0..i. (tps :#: 0 + Suc i, tps :#: 1 + Suc i)) [0..i. (tps :#: 0, tps :#: 1 + Suc i)) [0..i. (tps :#: 0, tps :#: 1 + Suc i)) [0.. 0" and "j < length tps" and "es = [(tps :#: 0, tps :#: 1)]" and "tps' = tps[j := tps ! j |:=| h]" shows "traces (tm_write j h) tps es tps'" using assms execute_tm_write tm_write_def by (intro tracesI) (auto simp add: nth_list_update) corollary traces_tm_write_1I: assumes "1 < length tps" and "es = [(tps :#: 0, tps :#: 1)]" and "tps' = tps[1 := tps ! 1 |:=| h]" shows "traces (tm_write 1 h) tps es tps'" using assms traces_tm_writeI by simp corollary traces_tm_write_ge2I: assumes "j \ 2" and "j < length tps" and "es = [(tps :#: 0, tps :#: 1)]" and "tps' = tps[j := tps ! j |:=| h]" shows "traces (tm_write j h) tps es tps'" using assms traces_tm_writeI by simp lemma execute_tm_write_manyI: assumes "0 \ J" and "\j\J. j < k" and "k \ 2" and "length tps = k" and "length tps' = k" and "\j. j \ J \ tps' ! j = tps ! j |:=| h" and "\j. j < k \ j \ J \ tps' ! j = tps ! j" shows "execute (tm_write_many J h) (0, tps) 1 = (1, tps')" proof - have "tps' = map (\j. if j \ J then tps ! j |:=| h else tps ! j) [0.. J" and "\j\J. j < k" and "k \ 2" and "length tps = k" and "length tps' = k" and "\j. j \ J \ tps' ! j = tps ! j |:=| h" and "\j. j < k \ j \ J \ tps' ! j = tps ! j" and "es = [(tps :#: 0, tps :#: 1)]" shows "traces (tm_write_many J h) tps es tps'" proof show "execute (tm_write_many J h) (0, tps) (length es) = (length (tm_write_many J h), tps')" using execute_tm_write_manyI[OF assms(1-7)] tm_write_many_def assms(8) by simp show "\i. i < length es \ fst (execute (tm_write_many J h) (0, tps) i) < length (tm_write_many J h)" using assms(8) tm_write_many_def by simp show "\i. i < length es \ snd (execute (tm_write_many J h) (0, tps) (Suc i)) :#: 0 = fst (es ! i) \ snd (execute (tm_write_many J h) (0, tps) (Suc i)) :#: 1 = snd (es ! i)" using execute_tm_write_manyI[OF assms(1-7)] tm_write_many_def assms(3,6,7,8) by (metis One_nat_def Suc_1 Suc_lessD fst_conv length_Cons less_Suc0 less_eq_Suc_le list.size(3) nth_Cons_0 snd_conv) qed lemma traces_tm_write_repeat_1I: assumes "1 < length tps" and "es = map (\i. (tps :#: 0, tps :#: 1 + Suc i)) [0..i. i < length es \ fst (execute ?M (0, tps) i) < length ?M" using assms execute_tm_write_repeat tm_write_repeat_def by simp show "execute ?M (0, tps) (Suc i) <#> 0 = fst (es ! i) \ execute ?M (0, tps) (Suc i) <#> 1 = snd (es ! i)" if "i < length es" for i proof - have "Suc i \ m" using assms \length es = m\ that by linarith then have "execute ?M (0, tps) (Suc i) = (Suc i, tps[1 := overwrite (tps ! 1) h (Suc i)])" using that execute_tm_write_repeat assms by blast then show ?thesis using overwrite_def assms(1,2) that by simp qed qed subsection \Memorizing in states\ text \ We need some results for the traces of ``cartesian'' machines used for the memorizing-in-states technique introduced in Section~\ref{s:tm-memorizing}. \ lemma cartesian_trace: assumes "turing_machine (Suc k) G M" and "immobile M k (Suc k)" and "M' = cartesian M G" and "k \ 2" and "\i 0" using assms(1) turing_machine_sequential_def by (simp add: turing_machine_def) show "fst (execute M' (start_config k zs) i) < length M'" if "i < length es" for i proof (rule ccontr) assume "\ fst (execute M' (start_config k zs) i) < length M'" then have "fst (execute M' (start_config k zs) i) \ length M'" by simp then have "fst (execute M' (start_config k zs) i) = length M'" using assms(1,3) cartesian_tm' by (metis (no_types, lifting) Suc_1 Suc_le_D assms(4) start_config_def start_config_length le_add2 le_add_same_cancel2 le_antisym less_Suc_eq_0_disj prod.sel(1) turing_machine_execute_states) then have "fst (squish G (length M) (execute M (start_config (Suc k) zs) i)) = G * length M" using assms cartesian_execute_start_config len by simp moreover have "fst (execute M (start_config (Suc k) zs) i) \ length M" using assms(1) assms(6) that trace_def by auto ultimately have "fst (execute M (start_config (Suc k) zs) i) = length M" using squish_halt_state `0 < G` by simp then show False using that assms(6) trace_def by auto qed show "execute M' (start_config k zs) (Suc i) <#> 0 = fst (es ! i) \ execute M' (start_config k zs) (Suc i) <#> 1 = snd (es ! i)" if "i < length es" for i proof (rule ccontr) assume a: "\ (snd (execute M' (start_config k zs) (Suc i)) :#: 0 = fst (es ! i) \ snd (execute M' (start_config k zs) (Suc i)) :#: 1 = snd (es ! i))" have *: "execute M' (start_config k zs) (Suc i) = squish G (length M) (execute M (start_config (Suc k) zs) (Suc i))" using assms cartesian_execute_start_config by blast then have "execute M' (start_config k zs) (Suc i) <#> 0 = squish G (length M) (execute M (start_config (Suc k) zs) (Suc i)) <#> 0" by simp also have "... = execute M (start_config (Suc k) zs) (Suc i) <#> 0" using squish_head_pos assms execute_num_tapes start_config_length le_imp_less_Suc zero_less_Suc by presburger also have "... = fst (es ! i)" using that assms trace_def by simp finally have fst: "execute M' (start_config k zs) (Suc i) <#> 0 = fst (es ! i)" . from * have "execute M' (start_config k zs) (Suc i) <#> 1 = squish G (length M) (execute M (start_config (Suc k) zs) (Suc i)) <#> 1" by simp also have "... = execute M (start_config (Suc k) zs) (Suc i) <#> 1" using squish_head_pos assms execute_num_tapes start_config_length le_imp_less_Suc zero_less_Suc by presburger also have "... = snd (es ! i)" using that assms trace_def by simp finally have "execute M' (start_config k zs) (Suc i) <#> 1 = snd (es ! i)" . then show False using fst a by simp qed qed lemma cartesian_traces: assumes "turing_machine (Suc k) G M" and "immobile M k (Suc k)" and "M' = cartesian M G" and "k \ 2" and "\ij. j < k \ immobile M (j + 2) (k + 2)" and "\i 2" by simp moreover have "traces M (snd (start_config (Suc (k + 2)) zs)) es tps" using Suc by simp ultimately have *: "traces ?M (snd (start_config (k + 2) zs)) es (butlast tps)" using assms(3) cartesian_traces by simp have "turing_machine (k + 2) G ?M" using \2 \ k + 2\ \turing_machine (Suc (k + 2)) G M\ cartesian_tm' by blast moreover have "\j. j < k \ immobile ?M (j + 2) (k + 2)" using cartesian_immobile Suc by simp ultimately have "traces (icartesian k ?M G) (snd (start_config 2 zs)) es (take 2 (butlast tps))" using * Suc by simp moreover have "take 2 (butlast tps) = take 2 tps" proof - have "length tps = Suc k + 2" using start_config_length traces_tapes_length Suc by (metis (mono_tags, lifting) add_gr_0 zero_less_Suc) then show ?thesis by (simp add: take_butlast) qed ultimately show ?case by simp qed end \ No newline at end of file diff --git a/thys/Cook_Levin/Oblivious_2_Tape.thy b/thys/Cook_Levin/Oblivious_2_Tape.thy --- a/thys/Cook_Levin/Oblivious_2_Tape.thy +++ b/thys/Cook_Levin/Oblivious_2_Tape.thy @@ -1,8281 +1,8279 @@ section \Existence of two-tape oblivious Turing machines\label{s:oblivious-two-tape}\ theory Oblivious_2_Tape imports Oblivious_Polynomial NP begin text \ In this section we show that for every polynomial-time multi-tape Turing machine there is a polynomial-time two-tape oblivious Turing machine that computes the same function and halts with its output tape head in cell number~1. Consider a $k$-tape Turing machine $M$ with polynomially bounded running time $T$. We construct a two-tape oblivious Turing machine $S$ simulating $M$ as follows. Lemma @{thm [source] polystructor} from the previous section provides us with a polynomial-time two-tape oblivious TM and a function $f(n) \geq T(n)$ such that the TM outputs $\mathbf{1}^{f(n)}$ for all inputs of length $n$. Executing this TM is the first thing our simulator does. The $f(n)$ symbols~@{text \} mark the space $S$ is going to use. Every cell $i=0, \dots, f(n) - 1$ of this space is to store a symbol that encodes a $(2k + 2)$-tuple consisting of: \begin{itemize} \item $k$ symbols from $M$'s alphabet representing the contents of all the $i$-th cells on the $k$ tapes of $M$; \item $k$ flags (called ``head flags'') signalling which of the $k$ tape heads of $M$ is in cell $i$; \item a flag (called ``counter flag'') initialized with 0; \item a flag (called ``start flag'') signalling whether $i = 0$. \end{itemize} Together the counter flags are a unary counter from 0 to $f(n)$. They are toggled from left to right. The start flags never change. The symbols and the head flags represent the $k$ tapes of $M$ at some step of the execution. By choice of $f$ the TM $M$ cannot use more than $f(n)$ cells on any tape. So the space marked with @{text \} symbols on the simulator's output tape suffices. Next the simulator initializes the space of @{text \} symbols with code symbols representing the start configuration of $M$ for the input given to the simulator. Then the main loop of the simulation performs $f(n)$ iterations. In each iteration $S$ performs one step of $M$'s computation. In order to do this it performs several left-to-right and right-to-left sweeps over all the $f(n)$ cells in the formatted section of the output tape. A sweep will move the output tape head one cell right (respectively left) in each step. In this way the simulator's head positions at any time will only depend on $f(n)$, and hence on $n$. Thus the machine will be oblivious. Moreover $f(n) \ge T(n)$, and so $M$ will be in the halting state after $f(n)$ iterations of the simulation. Counting the iterations to $f(n)$ is achieved via the counter flags. Finally the simulator extracts from each code symbol the symbol corresponding to $M$'s output tape, thus reconstructing the output of $M$ on the simulator's output tape. Thanks to the start flags, the simulator can easily locate the leftmost cell and put the output tape head one to the right of it, as required. The construction heavily uses the memorization-in-states technique (see Chapter~\ref{s:tm-memorizing}). At first the machine features $2k + 1$ memorization tapes in addition to the input tape and output tape. The purpose of those tapes is to store $M$'s state and the symbols under the tape heads of $M$ in the currently simulated step of $M$'s execution, as well as the $k$ symbols to be written write and head movements to be executed by the simulator. \ text \ The next predicate expresses that a Turing machine halts within a time bound depending on the length of the input. We did not have a need for this fairly basic predicate yet, because so far we were always interested in the halting configuration, too, and so the predicate @{const computes_in_time} sufficed. \null \ definition time_bound :: "machine \ nat \ (nat \ nat) \ bool" where "time_bound M k T \ \zs. bit_symbols zs \ fst (execute M (start_config k zs) (T (length zs))) = length M" lemma time_bound_ge: assumes "time_bound M k T" and "\n. fmt n \ T n" shows "time_bound M k fmt" using time_bound_def assms by (metis execute_after_halting_ge) text \ The time bound also bounds the position of all the tape heads. \ lemma head_pos_le_time_bound: assumes "turing_machine k G M" and "time_bound M k T" and "bit_symbols zs" and "j < k" shows "execute M (start_config k zs) t <#> j \ T (length zs)" using assms time_bound_def[of M k T] execute_after_halting_ge head_pos_le_time[OF assms(1,4)] by (metis (no_types, lifting) nat_le_linear) text \ The entire construction will take place in a locale that assumes a polynomial-time $k$-tape Turing machine $M$. \ locale two_tape = fixes M :: machine and k G :: nat and T :: "nat \ nat" assumes tm_M: "turing_machine k G M" and poly_T: "big_oh_poly T" and time_bound_T: "time_bound M k T" begin lemma k_ge_2: "k \ 2" using tm_M turing_machine_def by simp lemma G_ge_4: "G \ 4" using tm_M turing_machine_def by simp text \ The construction of the simulator relies on the formatted space on the output tape to be large enough to hold the input. The size of the formatted space depends on the time bound $T$, which might be less than the length of the input. To ensure that the formatted space is large enough we increase the time bound while keeping it polynomial. The new bound is $T'$: \ definition T' :: "nat \ nat" where "T' n \ T n + n" lemma poly_T': "big_oh_poly T'" proof - have "big_oh_poly (\n. n)" using big_oh_poly_poly[of 1] by simp then show ?thesis using T'_def big_oh_poly_sum[OF poly_T] by presburger qed lemma time_bound_T': "time_bound M k T'" using T'_def time_bound_ge[OF time_bound_T, of T'] by simp subsection \Encoding multiple tapes into one\ text \ The symbols on the output tape of the simulator are supposed to encode a $(2k + 2)$-tuple, where the first $k$ elements assume one of the values in $\{0, \dots, G - 1\}$, whereas the other $k + 2$ are flags with two possible values only. For uniformity we define an encoding where all elements range over $G$ values and that works for tuples of every length. \ fun encode :: "nat list \ nat" where "encode [] = 0" | "encode (x # xs) = x + G * encode xs" text \ For every $m \in \nat$, the enocoding is a bijective mapping from $\{0, \dots, G - 1\}^m$ to $\{0, \dots, G^m - 1\}$. \ lemma encode_surj: assumes "n < G ^ m" shows "\xs. length xs = m \ (\x\set xs. x < G) \ encode xs = n" using assms proof (induction m arbitrary: n) case 0 then show ?case by simp next case (Suc m) define q where "q = n div G" define r where "r = n mod G" have "q < G ^ m" using Suc q_def by (auto simp add: less_mult_imp_div_less power_commutes) then obtain xs' where xs': "length xs' = m" "\x\set xs'. x < G" "encode xs' = q" using Suc by auto have "r < G" using r_def G_ge_4 by simp have "encode (r # xs') = n" using xs'(3) q_def r_def G_ge_4 by simp moreover have "\x\set (r # xs'). x < G" using xs'(2) `r < G` by simp moreover have "length (r # xs') = Suc m" using xs'(1) by simp ultimately show ?case by blast qed lemma encode_inj: assumes "\x\set xs. x < G" and "length xs = m" and "\y\set ys. y < G" and "length ys = m" and "encode xs = encode ys" shows "xs = ys" using assms proof (induction m arbitrary: xs ys) case 0 then show ?case by simp next case (Suc m) obtain x xs' y ys' where xy: "xs = x # xs'" "ys = y # ys'" using Suc by (metis Suc_length_conv) then have len: "length xs' = m" "length ys' = m" using Suc by simp_all have *: "x + G * encode xs' = y + G * encode ys'" using Suc xy by simp moreover have "(x + G * encode xs') mod G = x" by (simp add: Suc.prems(1) xy(1)) moreover have "(y + G * encode ys') mod G = y" by (simp add: Suc.prems(3) xy(2)) ultimately have "x = y" by simp with * have "G * encode xs' = G * encode ys'" by simp then have "encode xs' = encode ys'" using G_ge_4 by simp with len Suc xy have "xs' = ys'" by simp then show ?case using `x = y` xy by simp qed lemma encode_less: assumes "\x\set xs. x < G" shows "encode xs < G ^ (length xs)" using assms proof (induction xs) case Nil then show ?case by simp next case (Cons x xs) then have "x < G" by simp have "encode (x # xs) = x + G * encode xs" by simp also have "... \ x + G * (G ^ (length xs) - 1)" using Cons by simp also have "... = x + G * G ^ (length xs) - G" by (simp add: right_diff_distrib') also have "... < G * G ^ (length xs)" using `x < G` less_imp_Suc_add by fastforce also have "... = G ^ (Suc (length xs))" by simp finally have "encode (x # xs) < G ^ (length (x # xs))" by simp then show ?case . qed text \ Decoding a number into an $m$-tuple of numbers is then a well-defined operation. \ definition decode :: "nat \ nat \ nat list" where "decode m n \ THE xs. encode xs = n \ length xs = m \ (\x\set xs. x < G)" lemma decode: assumes "n < G ^ m" shows encode_decode: "encode (decode m n) = n" and length_decode: "length (decode m n) = m" and decode_less_G: "\x\set (decode m n). x < G" proof - have "\xs. length xs = m \ (\x\set xs. x < G) \ encode xs = n" using assms encode_surj by simp then have *: "\!xs. encode xs = n \ length xs = m \ (\x\set xs. x < G)" using encode_inj by auto let ?xs = "decode m n" let ?P = "\xs. encode xs = n \ length xs = m \ (\x\set xs. x < G)" have "encode ?xs = n \ length ?xs = m \ (\x\set ?xs. x < G)" using * theI'[of ?P] decode_def by simp then show "encode (decode m n) = n" and "length (decode m n) = m" and "\x\set (decode m n). x < G" by simp_all qed lemma decode_encode: "\x\set xs. x < G \ decode (length xs) (encode xs) = xs" proof - fix xs :: "nat list" assume a: "\x\set xs. x < G" then have "encode xs < G ^ (length xs)" using encode_less by simp show "decode (length xs) (encode xs) = xs" unfolding decode_def proof (rule the_equality) show "encode xs = encode xs \ length xs = length xs \ (\x\set xs. x < G)" using a by simp show "\ys. encode ys = encode xs \ length ys = length xs \ (\x\set ys. x < G) \ ys = xs" using a encode_inj by simp qed qed text \ The simulator will access and update components of the encoded symbol. \ definition encode_nth :: "nat \ nat \ nat \ nat" where "encode_nth m n j \ decode m n ! j" definition encode_upd :: "nat \ nat \ nat \ nat \ nat" where "encode_upd m n j x \ encode ((decode m n) [j:=x])" lemma encode_nth_less: assumes "n < G ^ m" and "j < m" shows "encode_nth m n j < G" using assms encode_nth_def decode_less_G length_decode by simp text \ For the symbols the simulator actually uses, we fix $m = 2k + 2$ and reserve the symbols $\triangleright$ and $\Box$, effectively shifting the symbols by two. We call the symbols $\{2, \dots, G^{2k+2} + 2\}$ ``code symbols''. \ definition enc :: "symbol list \ symbol" where "enc xs \ encode xs + 2" definition dec :: "symbol \ symbol list" where "dec n \ decode (2 * k + 2) (n - 2)" lemma dec: assumes "n > 1" and "n < G ^ (2 * k + 2) + 2" shows enc_dec: "enc (dec n) = n" and length_dec: "length (dec n) = 2 * k + 2" and dec_less_G: "\x\set (dec n). x < G" proof - show "enc (dec n) = n" using enc_def dec_def encode_decode assms by simp show "length (dec n) = 2 * k + 2" using enc_def dec_def length_decode assms by simp show "\x\set (dec n). x < G" using enc_def dec_def decode_less_G assms by (metis Suc_leI less_diff_conv2 one_add_one plus_1_eq_Suc) qed lemma dec_enc: "\x\set xs. x < G \ length xs = 2 * k + 2 \ dec (enc xs) = xs" unfolding enc_def dec_def using decode_encode by fastforce definition enc_nth :: "nat \ nat \ nat" where "enc_nth n j \ dec n ! j" lemma enc_nth: "\x\set xs. x < G \ length xs = 2 * k + 2 \ enc_nth (enc xs) j = xs ! j" unfolding enc_nth_def by (simp add: dec_enc) lemma enc_nth_dec: assumes "n > 1" and "n < G ^ (2 * k + 2) + 2" shows "enc_nth n j = (dec n) ! j" using assms enc_nth dec by metis abbreviation is_code :: "nat \ bool" where "is_code n \ n < G ^ (2 * k + 2) + 2 \ 1 < n" lemma enc_nth_less: assumes "is_code n" and "j < 2 * k + 2" shows "enc_nth n j < G" using assms enc_nth_def by (metis dec_less_G in_set_conv_nth length_dec) lemma enc_less: "\x\set xs. x < G \ length xs = 2 * k + 2 \ enc xs < G ^ (2 * k + 2) + 2" using encode_less enc_def by fastforce definition enc_upd :: "nat \ nat \ nat \ nat" where "enc_upd n j x \ enc ((dec n) [j:=x])" lemma enc_upd_is_code: assumes "is_code n" and "j < 2 * k + 2" and "x < G" shows "is_code (enc_upd n j x)" proof - let ?ys = "(dec n) [j:=x]" have "\h\set (dec n). h < G" using assms(1,2) dec_less_G by auto then have "\h\set ?ys. h < G" using assms(3) by (metis in_set_conv_nth length_list_update nth_list_update_eq nth_list_update_neq) moreover have "length ?ys = 2 * k + 2" using assms length_dec by simp ultimately have "enc ?ys < G ^ (2 * k + 2) + 2" using enc_less by simp then show ?thesis using enc_upd_def enc_def by simp qed text \ The code symbols require the simulator to have an alphabet of at least size $G^{2k + 2} + 2$. On top of that we want to store on a memorization tape the state that $M$ is in. So the alphabet must have at least @{term "length M + 1"} symbols. Both conditions are met by the simulator having an alphabet of size $G'$: \ definition G' :: nat where "G' \ G ^ (2 * k + 2) + 2 + length M" lemma G'_ge_6: "G' \ 6" proof - have "4 ^ 2 > (5::nat)" by simp then have "G ^ 2 > (5::nat)" using G_ge_4 less_le_trans power2_nat_le_eq_le by blast then have "G ^ (2 * k + 2) > (5::nat)" using k_ge_2 G_ge_4 by (metis (no_types, opaque_lifting) dec_induct le_add2 order_less_le_subst1 pow_mono zero_less_Suc zero_less_numeral) then show ?thesis using G'_def by simp qed corollary G'_ge: "G' \ 4" "G' \ 5" using G'_ge_6 by simp_all lemma G'_ge_G: "G' \ G" proof - have "G ^ 2 \ G" by (simp add: power2_nat_le_imp_le) then have "G ^ (2 * k + 2) \ G" by simp then show ?thesis using G'_def by linarith qed corollary enc_less_G': "\x\set xs. x < G \ length xs = 2 * k + 2 \ enc xs < G'" using enc_less G'_def by fastforce lemma enc_greater: "enc xs > 1" using enc_def by simp subsection \Construction of the simulator Turing machine\ text \ The simulator is a sequence of three Turing machines: The ``formatter'', which initializes the output tape; the loop, which simulates the TM $M$ for polynomially many steps; and a cleanup TM, which makes the output tape look like the output tape of $M$. All these machines are discussed in order in the following subsections. The simulator will start with $2k + 1$ memorization tapes for a total of $2k + 3$ tapes. The simulation action will take place on the output tape. \ subsubsection \Initializing the simulator's tapes\ text \ The function @{const T'} is polynomially bounded and therefore there is a polynomial-time two-tape oblivious Turing machine that outputs at least $T'(n)$ symbols~@{text \} on an input of length $n$, as we have proven in the previous section (lemma~@{thm [source] polystructor}). We now obtain such a Turing machine together with its running time bound and trace. This TM will be one of our blocks for building the simulator TM. We will call it the ``formatter''. \null \ definition fmt_es_pM :: "(nat \ nat) \ (nat \ (nat \ nat) list) \ machine" where "fmt_es_pM \ SOME tec. turing_machine 2 G' (snd (snd tec)) \ big_oh_poly (\n. length ((fst (snd tec)) n)) \ big_oh_poly (fst tec) \ (\n. T' n \ (fst tec) n) \ (\zs. proper_symbols zs \ traces (snd (snd tec)) (start_tapes_2 zs) ((fst (snd tec)) (length zs)) (one_tapes_2 zs ((fst tec) (length zs))))" lemma polystructor': fixes GG :: nat and g :: "nat \ nat" assumes "big_oh_poly g" and "GG \ 5" shows "\f_es_M. turing_machine 2 GG (snd (snd f_es_M)) \ big_oh_poly (\n. length ((fst (snd f_es_M)) n)) \ big_oh_poly (fst f_es_M) \ (\n. g n \ (fst f_es_M) n) \ (\zs. proper_symbols zs \ traces (snd (snd f_es_M)) (start_tapes_2 zs) ((fst (snd f_es_M)) (length zs)) (one_tapes_2 zs ((fst f_es_M) (length zs))))" using polystructor[OF assms] by auto lemma fmt_es_pM: "turing_machine 2 G' (snd (snd fmt_es_pM)) \ big_oh_poly (\n. length ((fst (snd fmt_es_pM)) n)) \ big_oh_poly (fst fmt_es_pM) \ (\n. T' n \ (fst fmt_es_pM) n) \ (\zs. proper_symbols zs \ traces (snd (snd fmt_es_pM)) (start_tapes_2 zs) ((fst (snd fmt_es_pM)) (length zs)) (one_tapes_2 zs ((fst fmt_es_pM) (length zs))))" using fmt_es_pM_def polystructor'[OF poly_T' G'_ge(2)] someI_ex[of "\tec. turing_machine 2 G' (snd (snd tec)) \ big_oh_poly (\n. length ((fst (snd tec)) n)) \ big_oh_poly (fst tec) \ (\n. T' n \ (fst tec) n) \ (\zs. proper_symbols zs \ traces (snd (snd tec)) (start_tapes_2 zs) ((fst (snd tec)) (length zs)) (one_tapes_2 zs ((fst tec) (length zs))))"] by simp definition fmt :: "nat \ nat" where "fmt \ fst fmt_es_pM" definition es_fmt :: "nat \ (nat \ nat) list" where "es_fmt \ fst (snd fmt_es_pM)" definition tm_fmt :: "machine" where "tm_fmt \ snd (snd fmt_es_pM)" text \ The formatter TM is @{const tm_fmt}, the number of @{text \} symbols written to the output tape on input of length $n$ is @{term "fmt n"}, and the trace on inputs of length $n$ is @{term "es_fmt n"}. \ lemma fmt: "turing_machine 2 G' tm_fmt" "big_oh_poly (\n. length (es_fmt n))" "big_oh_poly fmt" "\n. T' n \ fmt n" "\zs. proper_symbols zs \ traces tm_fmt (start_tapes_2 zs) (es_fmt (length zs)) (one_tapes_2 zs (fmt (length zs)))" unfolding fmt_def es_fmt_def tm_fmt_def using fmt_es_pM by simp_all lemma fmt_ge_n: "fmt n \ n" using fmt(4) T'_def by (metis dual_order.strict_trans2 le_less_linear not_add_less2) text \ The formatter is a two-tape TM. The first incarnation of the simulator will have two tapes and $2k + 1$ memorization tapes. So for now we formally need to extend the formatter to $2k + 3$ tapes: \ definition "tm1 \ append_tapes 2 (2 * k + 3) tm_fmt" lemma tm1_tm: "turing_machine (2 * k + 3) G' tm1" unfolding tm1_def using append_tapes_tm by (simp add: fmt(1)) text \ Next we replace all non-blank symbols on the output tape by code symbols representing the tuple of $2k + 2$ zeros. \ definition "tm1_2 \ tm_const_until 1 1 {\} (enc (replicate k 0 @ replicate k 0 @ [0, 0]))" lemma tm1_2_tm: "turing_machine (2 * k + 3) G' tm1_2" unfolding tm1_2_def using G'_ge proof (intro tm_const_until_tm, auto) show "enc (replicate k 0 @ replicate k 0 @ [0, 0]) < G'" using G_ge_4 by (intro enc_less_G', auto) qed definition "tm2 \ tm1 ;; tm1_2" lemma tm2_tm: "turing_machine (2 * k + 3) G' tm2" unfolding tm2_def using tm1_tm tm1_2_tm by simp definition "tm3 \ tm2 ;; tm_start 1" lemma tm3_tm: "turing_machine (2 * k + 3) G' tm3" unfolding tm3_def using tm2_tm tm_start_tm G'_ge by simp text \ Back at the start symbol of the output tape, we replace it by the code symbol for the tuple $1^k1^k01$. The first $k$ ones represent that initially the $k$ tapes of $M$ have the start symbol (numerically 1) in the leftmost cell. The second run of $k$ ones represent that initially all $k$ tapes have their tape heads in the leftmost cell. The following 0 is the first bit of the unary counter, currently set to zero. The final flag~1 signals that this is the leftmost cell. Unlike the start symbols this signal flag cannot be overwritten by $M$. \ definition "tm4 \ tm3 ;; tm_write 1 (enc (replicate k 1 @ replicate k 1 @ [0, 1]))" lemma tm3_4_tm: "turing_machine (2 * k + 3) G' (tm_write 1 (enc (replicate k 1 @ replicate k 1 @ [0, 1])))" using G'_ge enc_less_G' G_ge_4 tm_write_tm by simp lemma tm4_tm: "turing_machine (2 * k + 3) G' tm4" unfolding tm4_def using tm3_tm tm3_4_tm by simp definition "tm5 \ tm4 ;; tm_right 1" lemma tm5_tm: "turing_machine (2 * k + 3) G' tm5" unfolding tm5_def using tm4_tm by (simp add: G'_ge(1) tm_right_tm) text \ So far the simulator's output tape encodes $k$ tapes that are empty but for the start symbols. To represent the start configuration of $M$, we need to copy the contents of the input tape to the output tape. The following TM moves the work head to the first blank while copying the input tape content. Here we exploit $T'(n) \geq n$, which implies that the formatted section is long enough to hold the input. \ definition "tm5_6 \ tm_trans_until 0 1 {0} (\h. enc (h mod G # replicate (k - 1) 0 @ replicate k 0 @ [0, 0]))" definition "tm6 \ tm5 ;; tm5_6" lemma tm5_6_tm: "turing_machine (2 * k + 3) G' tm5_6" unfolding tm5_6_def proof (rule tm_trans_until_tm, auto simp add: G'_ge(1) G_ge_4 k_ge_2 enc_less_G') show "\h. h < G' \ enc (h mod G # replicate (k - Suc 0) 0 @ replicate k 0 @ [0, 0]) < G'" using G_ge_4 k_ge_2 enc_less_G' by simp qed lemma tm6_tm: "turing_machine (2 * k + 3) G' tm6" unfolding tm6_def using tm5_tm tm5_6_tm by simp text \ Since we have overwritten the leftmost cell of the output tape with some code symbol, we cannot return to it using @{const tm_start}. But we can use @{const tm_left_until} with a special set of symbols: \ abbreviation StartSym :: "symbol set" where "StartSym \ {y. y < G ^ (2 * k + 2) + 2 \ y > 1 \ dec y ! (2 * k + 1) = 1}" abbreviation "tm_left_until1 \ tm_left_until StartSym 1" lemma tm_left_until1_tm: "turing_machine (2 * k + 3) G' tm_left_until1" using tm_left_until_tm G'_ge(1) k_ge_2 by simp definition "tm7 \ tm6 ;; tm_left_until1" lemma tm7_tm: "turing_machine (2 * k + 3) G' tm7" unfolding tm7_def using tm6_tm tm_left_until1_tm by simp text \ Tape number $2$ is meant to memorize $M$'s state during the simulation. Initially the state is the start state, that is, zero. \ definition "tm8 \ tm7 ;; tm_write 2 0" lemma tm8_tm: "turing_machine (2 * k + 3) G' tm8" unfolding tm8_def using tm7_tm tm_write_tm k_ge_2 G'_ge(1) by simp text \ We also initialize memorization tapes $3, \dots, 2k + 2$ to zero. This concludes the initialization of the simulator's tapes. \ definition "tm9 \ tm8 ;; tm_write_many {3..<2 * k + 3} 0" lemma tm9_tm: "turing_machine (2 * k + 3) G' tm9" unfolding tm9_def using tm8_tm tm_write_many_tm k_ge_2 G'_ge(1) by simp subsubsection \The loop\ text \ The core of the simulator is a loop whose body is executed @{term "fmt n"} many times. Each iteration simulates one step of the Turing machine $M$. For the analysis of the loop we describe the $2k + 3$ tapes in the form @{term "[a, b, c] @ map f1 [0.. lemma threeplus2k_2: assumes "3 \ j \ j < k + 3" shows "([a, b, c] @ map f1 [0.. j \ j < 2 * k + 3" shows "([a, b, c] @ map f1 [0.. To ensure the loop runs for @{term "fmt n"} iterations we increment the unary counter in the code symbols in each iteration. The loop terminates when there are no more code symbols with an unset counter flag. The TM that prepares the loop condition sweeps the output tape left-to-right searching for the first symbol that is either blank or has an unset counter flag. The condition then checks for which of the two cases occurred. This is the condition TM: \ definition "tmC \ tm_right_until 1 {y. y < G ^ (2 * k + 2) + 2 \ (y = 0 \ dec y ! (2 * k) = 0)}" lemma tmC_tm: "turing_machine (2 * k + 3) G' tmC" using tmC_def tm_right_until_tm using G'_ge(1) by simp text \ At the start of the iteration, the memorization tape~2 has the state of $M$, and all other memorization tapes contain $0$. The output tape head is at the leftmost code symbol with unset counter flag. The first order of business is to move back to the beginning of the output tape. \ definition "tmL1 \ tm_left_until1" lemma tmL1_tm: "turing_machine (2 * k + 3) G' tmL1" unfolding tmL1_def using tm6_tm tm_left_until1_tm by simp text \ Then the output tape head sweeps right until it encounters a blank. During the sweep the Turing machine checks for any set head flags, and if it finds the one for tape $j$ set, it memorizes the symbol for tape $j$ on tape $3 + k + j$. The next command performs this operation: \ definition cmdL2 :: command where "cmdL2 rs \ (if rs ! 1 = \ then (1, zip rs (replicate (2*k+3) Stay)) else (0, [(rs!0, Stay), (rs!1, Right), (rs!2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (if 1 < rs ! 1 \ rs ! 1 < G^(2*k+2)+2 \ enc_nth (rs!1) (k+j) = 1 then enc_nth (rs!1) j else rs!(3+k+j), Stay)) [0.. \" and "3 \ j \ j < k + 3" shows "cmdL2 rs [!] j = (rs ! j, Stay)" using cmdL2_def assms threeplus2k_2 by (metis (no_types, lifting) le_add_diff_inverse2 snd_conv) lemma cmdL2_at_4: assumes "rs ! 1 \ \" and "k + 3 \ j \ j < 2 * k + 3" shows "cmdL2 rs [!] j = (if 1 < rs ! 1 \ rs ! 1 < G^(2*k+2)+2 \ enc_nth (rs ! 1) (j-3) = 1 then enc_nth (rs ! 1) (j-k-3) else rs ! j, Stay)" unfolding cmdL2_def using assms threeplus2k_3[OF assms(2), of "(rs ! 0, Stay)"] by simp lemma cmdL2_at_4'': assumes "rs ! 1 \ \" and "k + 3 \ j \ j < 2 * k + 3" and "\ (1 < rs ! 1 \ rs ! 1 < G^(2*k+2)+2)" shows "cmdL2 rs [!] j = (rs ! j, Stay)" proof - have "cmdL2 rs [!] j = (if 1 < rs ! 1 \ rs ! 1 < G^(2*k+2)+2 \ enc_nth (rs!1) (j-3) = 1 then enc_nth (rs!1) (j-k-3) else rs!j, Stay)" using assms cmdL2_at_4 by blast then show ?thesis using assms(3) by auto qed lemma turing_command_cmdL2: "turing_command (2 * k + 3) 1 G' cmdL2" proof show "\gs. length gs = 2 * k + 3 \ length ([!!] cmdL2 gs) = length gs" unfolding cmdL2_def by simp show "\gs. length gs = 2 * k + 3 \ 0 < 2 * k + 3 \ cmdL2 gs [.] 0 = gs ! 0" unfolding cmdL2_def by simp show "cmdL2 gs [.] j < G'" if "length gs = 2 * k + 3" "\i. i < length gs \ gs ! i < G'" "j < length gs" for gs j proof (cases "gs ! 1 = 0") case True then have "cmdL2 gs = (1, zip gs (replicate (2*k+3) Stay))" unfolding cmdL2_def by simp have "cmdL2 gs [.] j = gs ! j" using that(1,3) by (simp add: \cmdL2 gs = (1, zip gs (replicate (2 * k + 3) Stay))\) then show ?thesis using that(2,3) by simp next case False consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using \j < length gs\ \length gs = 2 * k + 3\ by linarith then show ?thesis proof (cases) case 1 then show ?thesis by (simp add: cmdL2_def that(1) that(2)) next case 2 then show ?thesis unfolding cmdL2_def using False that by auto next case 3 then show ?thesis unfolding cmdL2_def using False that by auto next case 4 then have "snd (cmdL2 gs) ! j = (gs ! j, Stay)" unfolding cmdL2_def using False that threeplus2k_2[OF 4, of "(gs ! 0, Stay)"] by simp then show ?thesis using that by (simp add: \snd (cmdL2 gs) ! j = (gs ! j, Stay)\) next case 5 show ?thesis proof (cases "1 < gs ! 1 \ gs ! 1 < G ^ (2*k+2) + 2") case True then have *: "cmdL2 gs [.] j = (if enc_nth (gs ! 1) (j-3) = 1 then enc_nth (gs ! 1) (j-k-3) else gs ! j)" using 5 False by (simp add: cmdL2_at_4) let ?n = "gs ! 1" have len: "length (dec ?n) = 2 * k + 2" and less_G: "\x\set (dec ?n). x < G" using True length_dec dec_less_G by (simp, blast) have **: "enc_nth ?n (j-k-3) = dec ?n ! (j-k-3)" using enc_nth_dec True by simp then have "dec ?n ! (j-k-3) < G" using less_G 5 len by auto then have "dec ?n ! (j-k-3) < G'" using G'_ge_G by simp moreover have "gs ! j < G'" using that by simp ultimately show ?thesis using * ** by simp next case 6: False have "cmdL2 gs [!] j = (gs ! j, Stay)" using cmdL2_at_4''[OF False 5 6] by simp then show ?thesis using that by (simp add: \snd (cmdL2 gs) ! j = (gs ! j, Stay)\) qed qed qed show "\gs. length gs = 2 * k + 3 \ [*] (cmdL2 gs) \ 1" using cmdL2_def by simp qed definition "tmL1_2 \ [cmdL2]" lemma tmL1_2_tm: "turing_machine (2 * k + 3) G' tmL1_2" using tmL1_2_def using turing_command_cmdL2 G'_ge by auto definition "tmL2 \ tmL1 ;; tmL1_2" lemma tmL2_tm: "turing_machine (2 * k + 3) G' tmL2" by (simp add: tmL1_2_tm tmL1_tm tmL2_def) text \ The memorization tapes $3, \dots, 2 + k$ will store the head movements for tapes $0, \dots, k - 1$ of $M$. The directions are encoded as symbols thus: \ definition direction_to_symbol :: "direction \ symbol" where "direction_to_symbol m \ (case m of Left \ \ | Stay \ \ | Right \ \)" lemma direction_to_symbol_less: "direction_to_symbol m < 3" using direction_to_symbol_def by (cases m) simp_all text \ At this point in the iteration the memorization tapes $k + 3, \dots, 2k+2$ contain the symbols under the $k$ tape heads of $M$, and tape $2$ contains the state $M$ is in. Therefore all information is available to determine the actions $M$ is taking in the step currently simulated. The next command has the entire behavior of $M$ ``hard-coded'' and causes the actions to be stored on memorization tapes $3, \dots, 2k+2$: the output symbols on the tapes $k + 3, \dots, 2k + 2$, and the $k$ head movements on the tapes $3, \dots, k + 2$. The follow-up state will again be memorized on tape $2$. All this happens in one step of the simulator machine. \ definition cmdL3 :: command where "cmdL3 rs \ (1, [(rs ! 0, Stay), (rs ! 1, Stay), (if rs ! 2 < length M \ (\h\set (drop (k + 3) rs). h < G) then fst ((M ! (rs ! 2)) (drop (k + 3) rs)) else rs ! 2, Stay)] @ map (\j. (if rs ! 2 < length M \ (\h\set (drop (k + 3) rs). h < G) then direction_to_symbol ((M ! (rs ! 2)) (drop (k + 3) rs) [~] j) else 1, Stay)) [0..j. (if rs ! 2 < length M \ (\h\set (drop (k + 3) rs). h < G) then ((M ! (rs ! 2)) (drop (k + 3) rs) [.] j) else rs ! (k + 3 + j), Stay)) [0.. (\h\set (drop (k + 3) gs). h < G)" shows "cmdL3 gs [!] 2 = (fst ((M ! (gs ! 2)) (drop (k + 3) gs)), Stay)" using cmdL3_def assms by simp lemma cmdL3_at_2b: assumes "\ (gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G))" shows "cmdL3 gs [!] 2 = (gs ! 2, Stay)" using cmdL3_def assms by auto lemma cmdL3_at_3a: assumes "3 \ j \ j < k + 3" and "gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G)" shows "cmdL3 gs [!] j = (direction_to_symbol ((M ! (gs ! 2)) (drop (k + 3) gs) [~] (j - 3)), Stay)" using cmdL3_def assms(2) threeplus2k_2[OF assms(1), of "(gs ! 0, Stay)"] by simp lemma cmdL3_at_3b: assumes "3 \ j \ j < k + 3" and "\ (gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G))" shows "cmdL3 gs [!] j = (1, Stay)" using cmdL3_def assms(2) threeplus2k_2[OF assms(1), of "(gs ! 0, Stay)"] by auto lemma cmdL3_at_4a: assumes "k + 3 \ j \ j < 2 * k + 3" and "gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G)" shows "cmdL3 gs [!] j = ((M ! (gs ! 2)) (drop (k + 3) gs) [.] (j - k - 3), Stay)" using cmdL3_def assms(2) threeplus2k_3[OF assms(1), of "(gs ! 0, Stay)"] by simp lemma cmdL3_at_4b: assumes "k + 3 \ j \ j < 2 * k + 3" and "\ (gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G))" shows "cmdL3 gs [!] j = (gs ! j, Stay)" using assms cmdL3_def threeplus2k_3[OF assms(1), of "(gs ! 0, Stay)"] by auto lemma cmdL3_if_comm: assumes "length gs = 2 * k + 3" and "gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G)" shows "length ([!!] (M ! (gs ! 2)) (drop (k + 3) gs)) = k" and "\j. j < k \ (M ! (gs ! 2)) (drop (k + 3) gs) [.] j < G" proof - let ?rs = "drop (k + 3) gs" let ?cmd = "M ! (gs ! 2)" have *: "turing_command k (length M) G ?cmd" using assms(2) tm_M turing_machineD(3) by simp then show "length ([!!] ?cmd ?rs) = k" using turing_commandD(1) assms(1) by simp have "\i. i < length ?rs \ ?rs ! i < G" using assms(2) nth_mem by blast then show "\j. j < k \ (M ! (gs ! 2)) (drop (k + 3) gs) [.] j < G" using * turing_commandD(2) assms by simp qed lemma turing_command_cmdL3: "turing_command (2 * k + 3) 1 G' cmdL3" proof show "\gs. length gs = 2 * k + 3 \ length ([!!] cmdL3 gs) = length gs" using cmdL3_def by simp show "\gs. length gs = 2 * k + 3 \ 0 < 2 * k + 3 \ cmdL3 gs [.] 0 = gs ! 0" using cmdL3_def by simp show "cmdL3 gs [.] j < G'" if "length gs = 2 * k + 3" "\i. i < length gs \ gs ! i < G'" "j < length gs" for gs j proof - consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using \j < length gs\ \length gs = 2 * k + 3\ by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that cmdL3_def by simp next case 2 then show ?thesis using that cmdL3_def by simp next case 3 then show ?thesis proof (cases "gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G)") case True have "length (drop (k + 3) gs) = k" using that(1) by simp then have "fst ((M ! (gs ! 2)) (drop (k + 3) gs)) \ length M" using True turing_commandD(4) tm_M turing_machineD(3) by blast moreover have "cmdL3 gs [.] j = fst ((M ! (gs ! 2)) (drop (k + 3) gs))" using cmdL3_def True 3 by simp ultimately show ?thesis using G'_def by simp next case False then have "cmdL3 gs [.] j = gs ! 2" using cmdL3_def 3 by auto then show ?thesis using 3 that(2,3) by simp qed next case 4 then show ?thesis proof (cases "gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G)") case True then have "cmdL3 gs [.] j = direction_to_symbol ((M ! (gs ! 2)) (drop (k + 3) gs) [~] (j - 3))" using cmdL3_at_3a 4 by simp then have "cmdL3 gs [.] j < 3" using direction_to_symbol_less by simp then show ?thesis using G'_ge by simp next case False then show ?thesis using cmdL3_at_3b[OF 4] G'_ge by simp qed next case 5 then show ?thesis proof (cases "gs ! 2 < length M \ (\h\set (drop (k + 3) gs). h < G)") case True then have "cmdL3 gs [.] j = (M ! (gs ! 2)) (drop (k + 3) gs) [.] (j - k - 3)" using cmdL3_at_4a[OF 5] by simp then have "cmdL3 gs [.] j < G" using cmdL3_if_comm True 5 that(1) by auto then show ?thesis using G'_ge_G by simp next case False then have "cmdL3 gs [.] j = gs ! j" using cmdL3_at_4b[OF 5] by simp then show ?thesis using that by simp qed qed qed show "\gs. length gs = 2 * k + 3 \ [*] (cmdL3 gs) \ 1" using cmdL3_def by simp qed definition "tmL2_3 \ [cmdL3]" lemma tmL2_3_tm: "turing_machine (2 * k + 3) G' tmL2_3" unfolding tmL2_3_def using G'_ge(1) turing_command_cmdL3 by auto definition "tmL3 \ tmL2 ;; tmL2_3" lemma tmL3_tm: "turing_machine (2 * k + 3) G' tmL3" by (simp add: tmL2_3_tm tmL2_tm tmL3_def) text \ Next the output tape head sweeps left to the beginning of the tape, otherwise doing nothing. \ definition "tmL4 \ tmL3 ;; tm_left_until1" lemma tmL4_tm: "turing_machine (2 * k + 3) G' tmL4" using tmL3_tm tmL4_def tmL1_def tm_left_until1_tm by simp text \ The next four commands @{term cmdL4}, @{term cmdL5}, @{term cmdL6}, @{term cmdL7} are parameterized by $jj = 0, \dots, k - 1$. Their job is applying the write operation and head movement for tape $jj$ of $M$. The entire block of commands will then be executed $k$ times, once for each $jj$. The first of these commands sweeps right again and applies the write operation for tape $jj$, which is memorized on tape $3 + k + jj$. To this end it checks for head flags and updates the code symbol component $jj$ with the contents of tape $3+k+jj$ when the head flag for tape $jj$ is set. \ definition "cmdL5 jj rs \ if rs ! 1 = \ then (1, zip rs (replicate (2*k+3) Stay)) else (0, [(rs ! 0, Stay), (if is_code (rs ! 1) \ rs ! (3+k+jj) < G \ enc_nth (rs ! 1) (k+jj) = 1 then enc_upd (rs ! 1) jj (rs ! (3+k+jj)) else rs ! 1, Right), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. 0" shows "cmdL5 jj gs [!] 0 = (gs ! 0, Stay)" unfolding cmdL5_def using assms by simp lemma cmdL5_at_1: assumes "gs ! 1 \ 0" and "is_code (gs ! 1) \ gs ! (3+k+jj) < G \ enc_nth (gs!1) (k+jj) = 1" shows "cmdL5 jj gs [!] 1 = (enc_upd (gs!1) jj (gs!(3+k+jj)), Right)" unfolding cmdL5_def using assms by simp lemma cmdL5_at_1_else: assumes "gs ! 1 \ 0" and "\ (is_code (gs ! 1) \ gs ! (3+k+jj) < G \ enc_nth (gs!1) (k+jj) = 1)" shows "cmdL5 jj gs [!] 1 = (gs ! 1, Right)" unfolding cmdL5_def using assms by auto lemma cmdL5_at_2: assumes "gs ! 1 \ 0" shows "cmdL5 jj gs [!] 2 = (gs ! 2, Stay)" unfolding cmdL5_def using assms by simp lemma cmdL5_at_3: assumes "gs ! 1 \ 0" and "3 \ j \ j < 3 + k" shows "cmdL5 jj gs [!] j = (gs ! j, Stay)" unfolding cmdL5_def using assms threeplus2k_2[where ?a="(gs ! 0, Stay)"] by simp lemma cmdL5_at_4: assumes "gs ! 1 \ 0" and "3 + k \ j \ j < 2 * k + 3" shows "cmdL5 jj gs [!] j = (gs ! j, Stay)" unfolding cmdL5_def using assms threeplus2k_3[where ?a="(gs ! 0, Stay)"] by simp lemma turing_command_cmdL5: assumes "jj < k" shows "turing_command (2 * k + 3) 1 G' (cmdL5 jj)" proof show "length gs = 2 * k + 3 \ length ([!!] cmdL5 jj gs) = length gs" for gs unfolding cmdL5_def by (cases "gs!1=0") simp_all show goal2: "length gs = 2 * k + 3 \ 0 < 2 * k + 3 \ cmdL5 jj gs [.] 0 = gs ! 0" for gs unfolding cmdL5_def by (cases "gs ! 1=0") simp_all show "cmdL5 jj gs [.] j < G'" if "length gs = 2 * k + 3" "\i. i < length gs \ gs ! i < G'" "j < length gs" for gs j proof (cases "gs ! 1 = 0") case True then show ?thesis using that cmdL5_eq_0 by simp next case False consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `length gs = 2 * k + 3` `j < length gs` by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that goal2 by simp next case 2 show ?thesis proof (cases "1 < gs ! 1 \ gs ! 1 < G^(2*k+2)+2 \ gs ! (3+k+jj) < G \ enc_nth (gs ! 1) (k+jj) = 1") case True then have *: "cmdL5 jj gs [.] j = enc_upd (gs ! 1) jj (gs ! (3+k+jj))" using 2 cmdL5_at_1[OF False] by simp let ?n = "gs ! 1" let ?xs = "dec ?n" let ?ys = "(dec ?n) [jj:=gs!(3+k+jj)]" have "gs ! (3+k+jj) < G" using True by simp moreover have "\h\set (dec ?n). h < G" using True dec_less_G by auto ultimately have "\h\set ?ys. h < G" by (metis in_set_conv_nth length_list_update nth_list_update_eq nth_list_update_neq) moreover have "length ?ys = 2 * k + 2" using True length_dec by simp ultimately have "enc ?ys < G ^ (2 * k + 2) + 2" using enc_less by simp then show ?thesis using * by (simp add: enc_upd_def G'_def) next case b: False then show ?thesis using that cmdL5_at_1_else[OF False] 2 by simp qed next case 3 then show ?thesis using that cmdL5_at_2[OF False] by simp next case 4 then show ?thesis using that cmdL5_at_3[OF False] by simp next case 5 then show ?thesis using that cmdL5_at_4[OF False] by simp qed qed show "\gs. length gs = 2 * k + 3 \ [*] (cmdL5 jj gs) \ 1" using cmdL5_def by (metis (no_types, lifting) One_nat_def fst_conv le_eq_less_or_eq plus_1_eq_Suc trans_le_add2) qed definition tmL45 :: "nat \ machine" where "tmL45 jj \ [cmdL5 jj]" lemma tmL45_tm: assumes "jj < k" shows "turing_machine (2 * k + 3) G' (tmL45 jj)" using assms G'_ge turing_command_cmdL5 tmL45_def by auto text \ We move the output tape head one cell to the left. \ definition "tmL46 jj \ tmL45 jj ;; tm_left 1" lemma tmL46_tm: assumes "jj < k" shows "turing_machine (2 * k + 3) G' (tmL46 jj)" using assms G'_ge tm_left_tm tmL45_tm tmL46_def tmL45_def by simp text \ The next command sweeps left and applies the head movement for tape $jj$ if this is a movement to the left. To this end it checks for a set head flag in component $k + jj$ of the code symbol and clears it. It also memorizes that it just cleared the head flag by changing the symbol on memorization tape $3 + jj$ to the number $3$, which is not used to encode any actual head movement. In the next step of the sweep it will recognize this $3$ and set the head flag in component $k + jj$ of the code symbol. The net result is that the head flag for tape $jj$ has moved one cell to the left. \ abbreviation is_beginning :: "symbol \ bool" where "is_beginning y \ is_code y \ dec y ! (2 * k + 1) = 1" definition cmdL7 :: "nat \ command" where "cmdL7 jj rs \ (if is_beginning (rs ! 1) then 1 else 0, if rs ! (3 + jj) = \ \ enc_nth (rs ! 1) (k + jj) = 1 \ is_code (rs ! 1) \ \ is_beginning (rs ! 1) then [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 0, Left), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 3 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. is_code (rs ! 1) then [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 1, Left), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 0 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0..j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. gs ! (3 + jj) = 0 \ enc_nth (gs ! 1) (k + jj) = 1 \ is_code (gs ! 1) \ \ is_beginning (gs ! 1)" abbreviation "condition7b gs jj \ \ condition7a gs jj \ gs ! (3 + jj) = 3 \ is_code (gs ! 1)" abbreviation "condition7c gs jj \ \ condition7a gs jj \ \ condition7b gs jj" lemma turing_command_cmdL7: assumes "jj < k" shows "turing_command (2 * k + 3) 1 G' (cmdL7 jj)" proof show "length ([!!] cmdL7 jj gs) = length gs" if "length gs = 2 * k + 3" for gs proof - consider "condition7a gs jj" | "condition7b gs jj" | "condition7c gs jj" by blast then show ?thesis unfolding cmdL7_def using that by (cases) simp_all qed show goal2: "0 < 2 * k + 3 \ cmdL7 jj gs [.] 0 = gs ! 0" if "length gs = 2 * k + 3" for gs proof - consider "condition7a gs jj" | "condition7b gs jj" | "condition7c gs jj" by blast then show ?thesis unfolding cmdL7_def using that by (cases) simp_all qed show "cmdL7 jj gs [.] j < G'" if gs: "j < length gs" "length gs = 2 * k + 3" "\i. i < length gs \ gs ! i < G'" for gs j proof - consider "condition7a gs jj" | "condition7b gs jj" | "condition7c gs jj" by blast then show ?thesis proof (cases) case 1 then have *: "snd (cmdL7 jj gs) = [(gs ! 0, Stay), (enc_upd (gs ! 1) (k + jj) 0, Left), (gs ! 2, Stay)] @ (map (\j. (if j = jj then 3 else gs ! (j + 3), Stay)) [0..j. (gs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using gs by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that by (simp add: goal2) next case 2 then have "is_code (gs ! 1)" using 1 by blast moreover have "k + jj < 2 * k + 2" using assms by simp moreover have "0 < G" using G_ge_4 by simp ultimately have "is_code (enc_upd (gs ! 1) (k + jj) 0)" using enc_upd_is_code by simp then have "is_code (cmdL7 jj gs [.] j)" using * 2 by simp then show ?thesis using G'_ge_G G'_def by simp next case 3 then show ?thesis using * gs by simp next case 4 then show ?thesis using * gs G'_ge threeplus2k_2[where ?a="(gs ! 0, Stay)"] by simp next case 5 then show ?thesis using * gs G'_ge threeplus2k_3[where ?a="(gs ! 0, Stay)"] by simp qed next case case2: 2 then have *: "snd (cmdL7 jj gs) = [(gs ! 0, Stay), (enc_upd (gs ! 1) (k + jj) 1, Left), (gs ! 2, Stay)] @ (map (\j. (if j = jj then 0 else gs ! (j + 3), Stay)) [0..j. (gs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using gs by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that by (simp add: goal2) next case 2 then have "is_code (gs ! 1)" using case2 by simp moreover have "k + jj < 2 * k + 2" using assms by simp moreover have "1 < G" using G_ge_4 by simp ultimately have "is_code (enc_upd (gs ! 1) (k + jj) 1)" using enc_upd_is_code by simp then have "is_code (cmdL7 jj gs [.] j)" using * 2 by simp then show ?thesis using G'_ge_G G'_def by simp next case 3 then show ?thesis using * gs by simp next case 4 then show ?thesis using * gs G'_ge threeplus2k_2[where ?a="(gs ! 0, Stay)"] by simp next case 5 then show ?thesis using * gs G'_ge threeplus2k_3[where ?a="(gs ! 0, Stay)"] by simp qed next case case3: 3 then have *: "snd (cmdL7 jj gs) = [(gs ! 0, Stay), (gs ! 1, Left), (gs ! 2, Stay)] @ (map (\j. (gs ! (j + 3), Stay)) [0..j. (gs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using gs by linarith then show ?thesis using * gs G'_ge threeplus2k_2[where ?a="(gs ! 0, Stay)"] threeplus2k_3[where ?a="(gs ! 0, Stay)"] by (cases) simp_all qed qed show "\gs. length gs = 2 * k + 3 \ [*] (cmdL7 jj gs) \ 1" using cmdL7_def by simp qed definition tmL67 :: "nat \ machine" where "tmL67 jj \ [cmdL7 jj]" lemma tmL67_tm: assumes "jj < k" shows "turing_machine (2 * k + 3) G' (tmL67 jj)" using assms G'_ge tmL67_def turing_command_cmdL7 by auto definition "tmL47 jj \ tmL46 jj ;; tmL67 jj" lemma tmL47_tm: assumes "jj < k" shows "turing_machine (2 * k + 3) G' (tmL47 jj)" using assms G'_ge tm_left_tm tmL46_tm tmL47_def tmL67_tm by simp text \ Next we are sweeping right again and perform the head movement for tape $jj$ if this is a movement to the right. This works the same as the left movements in @{const cmdL7}. \ definition cmdL8 :: "nat \ command" where "cmdL8 jj rs \ (if rs ! 1 = \ then 1 else 0, if rs ! (3 + jj) = 2 \ enc_nth (rs ! 1) (k + jj) = 1 \ is_code (rs ! 1) then [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 0, Right), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 3 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. is_code (rs ! 1) then [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 1, Right), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 2 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0..j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0..j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. gs ! (3 + jj) = 2 \ enc_nth (gs ! 1) (k + jj) = 1 \ is_code (gs ! 1)" abbreviation "condition8b gs jj \ \ condition8a gs jj \ gs ! (3 + jj) = 3 \ is_code (gs ! 1)" abbreviation "condition8c gs jj \ \ condition8a gs jj \ \ condition8b gs jj \ gs ! 1 = 0" abbreviation "condition8d gs jj \ \ condition8a gs jj \ \ condition8b gs jj \ \ condition8c gs jj" lemma turing_command_cmdL8: assumes "jj < k" shows "turing_command (2 * k + 3) 1 G' (cmdL8 jj)" proof show "length ([!!] cmdL8 jj gs) = length gs" if "length gs = 2 * k + 3" for gs proof - consider "condition8a gs jj" | "condition8b gs jj" | "condition8c gs jj" | "condition8d gs jj" by blast then show ?thesis unfolding cmdL8_def using that by (cases) simp_all qed show goal2: "0 < 2 * k + 3 \ cmdL8 jj gs [.] 0 = gs ! 0" if "length gs = 2 * k + 3" for gs proof - consider "condition8a gs jj" | "condition8b gs jj" | "condition8c gs jj" | "condition8d gs jj" by blast then show ?thesis unfolding cmdL8_def using that by (cases) simp_all qed show "cmdL8 jj gs [.] j < G'" if gs: "j < length gs" "length gs = 2 * k + 3" "\i. i < length gs \ gs ! i < G'" for gs j proof - consider "condition8a gs jj" | "condition8b gs jj" | "condition8c gs jj" | "condition8d gs jj" by blast then show ?thesis proof (cases) case 1 then have *: "snd (cmdL8 jj gs) = [(gs ! 0, Stay), (enc_upd (gs ! 1) (k + jj) 0, Right), (gs ! 2, Stay)] @ (map (\j. (if j = jj then 3 else gs ! (j + 3), Stay)) [0..j. (gs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using gs by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that by (simp add: goal2) next case 2 then have "is_code (gs ! 1)" using 1 by blast moreover have "k + jj < 2 * k + 2" using assms by simp moreover have "0 < G" using G_ge_4 by simp ultimately have "is_code (enc_upd (gs ! 1) (k + jj) 0)" using enc_upd_is_code by simp then have "is_code (cmdL8 jj gs [.] j)" using * 2 by simp then show ?thesis using G'_ge_G G'_def by simp next case 3 then show ?thesis using * gs by simp next case 4 then show ?thesis using * gs G'_ge threeplus2k_2[where ?a="(gs ! 0, Stay)"] by simp next case 5 then show ?thesis using * gs G'_ge threeplus2k_3[where ?a="(gs ! 0, Stay)"] by simp qed next case case2: 2 then have *: "snd (cmdL8 jj gs) = [(gs ! 0, Stay), (enc_upd (gs ! 1) (k + jj) 1, Right), (gs ! 2, Stay)] @ (map (\j. (if j = jj then 2 else gs ! (j + 3), Stay)) [0..j. (gs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using gs by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that by (simp add: goal2) next case 2 then have "is_code (gs ! 1)" using case2 by simp moreover have "k + jj < 2 * k + 2" using assms by simp moreover have "1 < G" using G_ge_4 by simp ultimately have "is_code (enc_upd (gs ! 1) (k + jj) 1)" using enc_upd_is_code by simp then have "is_code (cmdL8 jj gs [.] j)" using * 2 by simp then show ?thesis using G'_ge_G G'_def by simp next case 3 then show ?thesis using * gs by simp next case 4 then show ?thesis using * gs G'_ge threeplus2k_2[where ?a="(gs ! 0, Stay)"] by simp next case 5 then show ?thesis using * gs G'_ge threeplus2k_3[where ?a="(gs ! 0, Stay)"] by simp qed next case 3 then have *: "snd (cmdL8 jj gs) = [(gs ! 0, Stay), (gs ! 1, Stay), (gs ! 2, Stay)] @ (map (\j. (gs ! (j + 3), Stay)) [0..j. (gs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using gs by linarith then show ?thesis using * gs G'_ge threeplus2k_2[where ?a="(gs ! 0, Stay)"] threeplus2k_3[where ?a="(gs ! 0, Stay)"] by (cases) simp_all next case 4 then have *: "snd (cmdL8 jj gs) = [(gs ! 0, Stay), (gs ! 1, Right), (gs ! 2, Stay)] @ (map (\j. (gs ! (j + 3), Stay)) [0..j. (gs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using gs by linarith then show ?thesis using * gs G'_ge threeplus2k_2[where ?a="(gs ! 0, Stay)"] threeplus2k_3[where ?a="(gs ! 0, Stay)"] by (cases) simp_all qed qed show "\gs. length gs = 2 * k + 3 \ [*] (cmdL8 jj gs) \ 1" using cmdL8_def by simp qed definition tmL78 :: "nat \ machine" where "tmL78 jj \ [cmdL8 jj]" lemma tmL78_tm: assumes "jj < k" shows "turing_machine (2 * k + 3) G' (tmL78 jj)" using assms G'_ge tmL78_def turing_command_cmdL8 by auto definition "tmL48 jj \ tmL47 jj ;; tmL78 jj" lemma tmL48_tm: assumes "jj < k" shows "turing_machine (2 * k + 3) G' (tmL48 jj)" using assms G'_ge tm_left_tm tmL47_tm tmL48_def tmL78_tm by simp text \ The last command in the command sequence is moving back to the beginning of the output tape. \ definition "tmL49 jj \ tmL48 jj ;; tm_left_until1" text \ The Turing machine @{term "tmL49 jj"} is then repeated for the parameters $jj = 0, \dots, k - 1$ in order to simulate the actions of $M$ on all tapes. \ lemma tmL49_tm: "jj < k \ turing_machine (2 * k + 3) G' (tmL49 jj)" using tmL48_tm tmL49_def tmL1_def tm_left_until1_tm by simp fun tmL49_upt :: "nat \ machine" where "tmL49_upt 0 = []" | "tmL49_upt (Suc j) = tmL49_upt j ;; tmL49 j" lemma tmL49_upt_tm: assumes "j \ k" shows "turing_machine (2 * k + 3) G' (tmL49_upt j)" using assms proof (induction j) case 0 then show ?case using G'_ge(1) turing_machine_def by simp next case (Suc j) then show ?case using assms tmL49_tm by simp qed definition "tmL9 \ tmL4 ;; tmL49_upt k" lemma tmL9_tm: "turing_machine (2 * k + 3) G' tmL9" unfolding tmL9_def using tmL49_upt_tm tmL4_tm by simp text \ At this point in the iteration we have completed one more step in the execution of $M$. We mark this be setting one more counter flag, namely the one in the leftmost code symbol where the flag is still unset. To find the first unset counter flag, we reuse @{term tmC}. \ definition "tmL10 \ tmL9 ;; tmC" lemma tmL10_tm: "turing_machine (2 * k + 3) G' tmL10" unfolding tmL10_def using tmL9_tm tmC_tm by simp text \ Then we set the counter flag, unless we have reached a blank symbol. \ definition cmdL11 :: command where "cmdL11 rs \ (1, [(rs ! 0, Stay), if is_code (rs ! 1) then (enc_upd (rs ! 1) (2 * k) 1, Stay) else (rs ! 1, Stay), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. length ([!!] cmdL11 gs) = length gs" for gs unfolding cmdL11_def by (cases "gs ! 1 = 0") simp_all show goal2: "length gs = 2 * k + 3 \ 0 < 2 * k + 3 \ cmdL11 gs [.] 0 = gs ! 0" for gs unfolding cmdL11_def by (cases "gs ! 1 = 0") simp_all show "cmdL11 gs [.] j < G'" if "length gs = 2 * k + 3" "\i. i < length gs \ gs ! i < G'" "j < length gs" for gs j proof - consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `length gs = 2 * k + 3` `j < length gs` by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that goal2 by simp next case 2 show ?thesis proof (cases "is_code (gs ! 1)") case True then have *: "cmdL11 gs [.] j = enc_upd (gs ! 1) (2 * k) 1" using 2 cmdL11_def by simp then have "is_code (cmdL11 gs [.] j)" using enc_upd_is_code[OF True] G_ge_4 by simp then show ?thesis using G'_def by simp next case False then show ?thesis using that cmdL11_def 2 by auto qed next case 3 then show ?thesis using that cmdL11_def by simp next case 4 then show ?thesis using that cmdL11_def threeplus2k_2[OF 4, of "(gs ! 0, Stay)"] by simp next case 5 then show ?thesis using that cmdL11_def threeplus2k_3[OF 5, of "(gs ! 0, Stay)"] by simp qed qed show "\gs. length gs = 2 * k + 3 \ [*] (cmdL11 gs) \ 1" using cmdL11_def by simp qed definition "tmL11 \ tmL10 ;; [cmdL11] " lemma tmL1011_tm: "turing_machine (2 * k + 3) G' [cmdL11]" using cmdL11_def turing_command_cmdL11 G'_ge by auto lemma tmL11_tm: "turing_machine (2 * k + 3) G' tmL11" using tmL11_def tmL1011_tm G'_ge tmL10_tm by simp text \ And we move back to the beginning of the output tape again. \ definition "tmL12 \ tmL11 ;; tm_left_until1" lemma tmL12_tm: "turing_machine (2 * k + 3) G' tmL12" using tmL11_tm tmL12_def tm_left_until1_tm by simp text \ Now, at the end of the iteration we set the memorization tapes $3, \dots, 2k+2$, that is, all but the one memorizing the state of $M$, to $0$. This makes for a simpler loop invariant than having the leftover symbols there. \ definition "tmL13 \ tmL12 ;; tm_write_many {3..<2 * k + 3} 0" lemma tmL13_tm: "turing_machine (2 * k + 3) G' tmL13" unfolding tmL13_def using tmL12_tm tm_write_many_tm k_ge_2 G'_ge(1) by simp text \ This is the entire loop. It terminates once there are no unset counter flags anymore. \ definition "tmLoop \ WHILE tmC ; \rs. rs ! 1 > \ DO tmL13 DONE" lemma tmLoop_tm: "turing_machine (2 * k + 3) G' tmLoop" using tmLoop_def turing_machine_loop_turing_machine tmC_tm tmL13_tm by simp definition "tm10 \ tm9 ;; tmLoop" lemma tm10_tm: "turing_machine (2 * k + 3) G' tm10" using tm10_def tmLoop_tm tm9_tm by simp subsubsection \Cleaning up the output\ text \ Now the simulation proper has ended, but the output tape does not yet look quite like the output tape of $M$. It remains to extract the component~$1$ from all the code symbols on the output tape. The simulator does this while sweeping left. Formally, ``extracting component~1'' means this: \ abbreviation ec1 :: "symbol \ symbol" where "ec1 h \ if is_code h then enc_nth h 1 else \" lemma ec1: "ec1 h < G'" if "h < G'" for h proof (cases "is_code h") case True then show ?thesis using enc_nth_less G'_ge_G by fastforce next case False then show ?thesis using that by auto qed definition "tm11 \ tm10 ;; tm_ltrans_until 1 1 StartSym ec1" lemma tm11_tm: "turing_machine (2 * k + 3) G' tm11" proof - have "turing_machine (2 * k + 3) G' (tm_ltrans_until 1 1 StartSym ec1)" using G'_ge ec1 by (intro tm_ltrans_until_tm) simp_all then show ?thesis using tm10_tm tm11_def by simp qed text \ The previous TM, @{term "tm_ltrans_until 1 1 StartSym ec1"}, halts as soon as it encounters a code symbol with the start flag set, without applying the extraction function. Applying the function to this final code symbol, which is at the leftmost cell of the tape, is the last step of the simulator machine. \ definition "tm12 \ tm11 ;; tm_rtrans 1 ec1" lemma tm12_tm: "turing_machine (2 * k + 3) G' tm12" unfolding tm12_def using tm_rtrans_tm tm11_tm ec1 G'_ge by simp subsection \Semantics of the Turing machine\ text \ This section establishes not only the configurations of the simulator but also the traces. For every Turing machine and command defined in the previous subsection, there will be a corresponding trace and a tape list representing the simulator's configuration after the command or TM has been applied. \ text \ For most of the time, the simulator's output tape will have no start symbol, and so the next definition will be more suited to describing it than the customary @{const contents}. \ definition contents' :: "symbol list \ (nat \ symbol)" where "contents' ys x \ if x < length ys then ys ! x else 0" lemma contents'_eqI: assumes "\x. x < length ys \ zs x = ys ! x" and "\x. x \ length ys \ zs x = 0" shows "zs = contents' ys" using contents'_def assms by auto lemma contents_contents': "\ys\ = contents' (1 # ys)" using contents_def contents'_def by auto lemma contents'_at_ge: assumes "i \ length ys" shows "contents' ys i = 0" using assms contents'_def by simp text \ In the following context @{term zs} represents the input for $M$ and hence for the simulator. \ context fixes zs :: "symbol list" assumes zs: "bit_symbols zs" begin lemma zs_less_G: "\i length zs" abbreviation "TT \ Suc (fmt n)" subsubsection \Initializing the simulator's tapes\ definition tps0 :: "tape list" where "tps0 \ [(\zs\, 0), (\[]\, 0)] @ replicate (2 * k + 1) (\\\)" lemma tps0_start_config: "start_config (2 * k + 3) zs = (0, tps0)" unfolding tps0_def contents_def onesie_def start_config_def by auto definition tps1 :: "tape list" where "tps1 \ [(\zs\, 1), (\replicate (fmt n) 3\, 1)] @ replicate (2 * k + 1) (\\\)" definition "es1 \ es_fmt n" lemma tm1: "traces tm1 tps0 es1 tps1" proof - let ?tps = "replicate (2 * k + 1) (\\\)" have 1: "tps0 = start_tapes_2 zs @ ?tps" using contents_def tps0_def start_tapes_2_def by auto have 2: "tps1 = one_tapes_2 zs (fmt n) @ ?tps" using contents_def tps1_def one_tapes_2_def by simp have "length (start_tapes_2 zs) = 2" using start_tapes_2_def by simp moreover have "traces tm_fmt (start_tapes_2 zs) (es_fmt n) (one_tapes_2 zs (fmt n))" using fmt zs by fastforce moreover have "turing_machine 2 G' tm_fmt" using fmt(1) . ultimately have "traces (append_tapes 2 (2 + length ?tps) tm_fmt) (start_tapes_2 zs @ ?tps) (es_fmt n) (one_tapes_2 zs (fmt n) @ ?tps)" using traces_append_tapes by blast then have "traces (append_tapes 2 (2 * k + 3) tm_fmt) (start_tapes_2 zs @ ?tps) (es_fmt n) (one_tapes_2 zs (fmt n) @ ?tps)" by (simp add: numeral_3_eq_3) then have "traces (append_tapes 2 (2 * k + 3) tm_fmt) tps0 (es_fmt n) tps1" using 1 2 by simp then show ?thesis using tm1_def es1_def by simp qed definition "es1_2 \ map (\i. (1, 1 + Suc i)) [0.. es1 @ es1_2" lemma len_es2: "length es2 = length (es_fmt n) + TT" using es2_def es1_def by (simp add: es1_2_def) definition tps2 :: "tape list" where "tps2 \ [(\zs\, 1), (\replicate (fmt n) (enc (replicate k 0 @ replicate k 0 @ [0, 0]))\, TT)] @ replicate (2 * k + 1) (\\\)" lemma tm2: "traces tm2 tps0 es2 tps2" unfolding tm2_def es2_def proof (rule traces_sequential) show "traces tm1 tps0 es1 tps1" using tm1 . show "traces tm1_2 tps1 es1_2 tps2" unfolding tm1_2_def es1_2_def proof (rule traces_tm_const_until_11I[where ?n="fmt n" and ?G=G']) show "1 < length tps1" using tps1_def by simp show "enc (replicate k 0 @ replicate k 0 @ [0, 0]) < G'" using G_ge_4 by (intro enc_less_G') simp_all show "rneigh (tps1 ! 1) {0} (fmt n)" using tps1_def contents_def by (intro rneighI) simp_all show "map (\i. (1, 1 + Suc i)) [0..i. (tps1 :#: 0, tps1 :#: 1 + Suc i)) [0.. es2 @ map (\i. (1, i)) (rev [0.. [(\zs\, 1), (\replicate (fmt n) (enc (replicate k 0 @ replicate k 0 @ [0, 0]))\, 0)] @ replicate (2 * k + 1) (\\\)" lemma tm3: "traces tm3 tps0 es3 tps3" unfolding tm3_def es3_def proof (rule traces_sequential) show "traces tm2 tps0 es2 tps2" using tm2 . show "traces (tm_start 1) tps2 (map (\i. (1, i)) (rev [0.. es3 @ [(1, 0)]" lemma len_es4: "length es4 = length (es_fmt n) + 2 * TT + 2" using es4_def es3_def len_es2 by simp definition tps4 :: "tape list" where "tps4 \ [(\zs\, 1), (contents' ((enc (replicate k 1 @ replicate k 1 @ [0, 1])) # replicate (fmt n) (enc (replicate k 0 @ replicate k 0 @ [0, 0]))), 0)] @ replicate (2 * k + 1) (\\\)" lemma tm4: "traces tm4 tps0 es4 tps4" unfolding tm4_def es4_def proof (rule traces_sequential) show "traces tm3 tps0 es3 tps3" using tm3 . show "traces (tm_write 1 (enc (replicate k 1 @ replicate k 1 @ [0, 1]))) tps3 [(1, 0)] tps4" proof (rule traces_tm_write_1I) show "1 < length tps3" using tps3_def by simp show "[(1, 0)] = [(tps3 :#: 0, tps3 :#: 1)]" using tps3_def by simp show "tps4 = tps3[1 := tps3 ! 1 |:=| enc (replicate k 1 @ replicate k 1 @ [0, 1])]" using tps3_def tps4_def contents'_def contents_contents' by auto qed qed definition "es5 \ es4 @ [(1, 1)]" lemma len_es5: "length es5 = length (es_fmt n) + 2 * TT + 3" using es5_def len_es4 by simp definition tps5 :: "tape list" where "tps5 \ [(\zs\, 1), (contents' ((enc (replicate k 1 @ replicate k 1 @ [0, 1])) # replicate (fmt n) (enc (replicate k 0 @ replicate k 0 @ [0, 0]))), 1)] @ replicate (2 * k + 1) (\\\)" lemma tm5: "traces tm5 tps0 es5 tps5" unfolding tm5_def es5_def proof (rule traces_sequential) show "traces tm4 tps0 es4 tps4" using tm4 . show "traces (tm_right 1) tps4 [(1, 1)] tps5" using tps4_def tps5_def by (intro traces_tm_right_1I) simp_all qed text \ Since the simulator simulates $M$ on @{term zs}, its tape contents are typically described in terms of configurations of $M$. So the following definition is actually more like an abbreviation. \ definition "exec t \ execute M (start_config k zs) t" lemma exec_pos_less_TT: assumes "j < k" shows "exec t <#> j < TT" proof - have "exec t <#> j \ T' n" using head_pos_le_time_bound[OF tm_M time_bound_T' zs assms] exec_def by simp then show ?thesis by (meson fmt(4) le_imp_less_Suc le_trans) qed lemma tps_ge_TT_0: assumes "i \ TT" shows "(exec t <:> 1) i = 0" proof (induction t) case 0 have "exec 0 = start_config k zs" using exec_def by simp then show ?case using start_config1 assms tm_M turing_machine_def by simp next case (Suc t) have *: "exec (Suc t) = exe M (exec t)" using exec_def by simp show ?case proof (cases "fst (exec t) \ length M") case True then have "exec (Suc t) = exec t" using * exe_def by simp then show ?thesis using Suc by simp next case False then have "exec (Suc t) <:> 1 = sem (M ! (fst (exec t))) (exec t) <:> 1" (is "_ = sem ?cmd ?cfg <:> 1") using exe_lt_length * by simp also have "... = fst (map (\(a, tp). act a tp) (zip (snd (?cmd (read (snd ?cfg)))) (snd ?cfg)) ! 1)" using sem' by simp also have "... = fst (act (snd (?cmd (read (snd ?cfg))) ! 1) (snd ?cfg ! 1))" (is "_ = fst (act ?h (snd ?cfg ! 1))") proof - have "||?cfg|| = k" using exec_def tm_M execute_num_tapes[OF tm_M] start_config_length turing_machine_def by simp moreover from this have "length (snd (?cmd (read (snd ?cfg)))) = k" by (metis False turing_commandD(1) linorder_not_less read_length tm_M turing_machineD(3)) moreover have "k > 1" using k_ge_2 by simp ultimately show ?thesis by simp qed finally have "exec (Suc t) <:> 1 = fst (act ?h (?cfg 1))" . moreover have "i \ snd (?cfg 1)" using assms by (metis Suc_1 exec_pos_less_TT lessI less_irrefl less_le_trans tm_M turing_machine_def) ultimately have "(exec (Suc t) <:> 1) i = fst (?cfg 1) i" using act_changes_at_most_pos by (metis prod.collapse) then show ?thesis using Suc.IH by simp qed qed text \ The next definition is central to how we describe the simulator's output tape. The basic idea is that it describes the tape during the simulation of the $t$-th step of executing $M$ on input @{term zs}. Recall that $TT$ is the length of the formatted part on the simulator's output tape. The $i$-th cell of the output tape contains: (1) $k$ symbols corresponding to the symbols in the $i$-th cell of the $k$ tapes of $M$ after $t$ steps; (2) $k$ flags indicating which of the $k$ tape heads is in position $i$; (3) a unary counter representing the number $t$; (4) a flag indicating whether $i = 0$. This is the situation at the beginning of the $t$-iteration of the simulator's main loop. During this iteration the tape changes slightly: some symbols and head positions may already represent the situation after $t+1$ steps of $M$, that is, the $t$-th step has been partially simulated already. To account for this, there is the @{term xs} parameter. It is meant to be set to a list of $k$ pairs. Let the $j$-th element of this list be $(a, b)$. on $M$'s tape $j$ has already been simulated. In other words, the output tape reflects the situation after $t + a$ steps. Likewise $b$ will be either None or 0 or 1. If it is 0 or 1, it means the flag represents the head position of tape $j$ after $t + b$ steps. If it is None, it means that all head flags for tape $k$ are currently zero, representing a ``tape without head''. This situation occurs every time the simulator has cleared the head flag representing the position after $t$ steps, bus has not set the flag for the head position after $t+1$ steps yet. Therefore, at the beginning of the $t$-t iteration of the simulator's loop @{term xs} consists of $k$ pairs $(0, 0)$. During the iteration every pair will eventually become $(0, 0)$. \ definition zip_cont :: "nat \ (nat \ nat option) list \ (nat \ symbol)" where "zip_cont t xs i \ if i < TT then enc (map (\j. (exec (t + fst (xs ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0.. Some auxiliary lemmas for accessing elements of lists of certain shape: \ lemma less_k_nth: "j < k \ (map f1 [0.. j \ j < 2 * k \ (map f1 [0.. j" "j < 2 * k" have b: "(map f1 [0..x\set (map (\j. (exec (t + fst (xs ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..x\set(?us @ ?vs @ [?a, ?b]). x < G") proof - let ?ys = "?us @ ?vs @ [?a, ?b]" let ?f1 = "(\j. (exec (t + fst (xs ! j)) <:> j) i)" let ?f2 = "(\j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0)" have "?ys ! j < G" if "j < 2 * k + 2" for j proof - consider "j < k" | "k \ j \ j < 2 * k" | "j = 2 * k" | "j = 2 * k + 1" using `j < 2 * k + 2` by linarith then show ?thesis proof (cases) case 1 then have "?us ! j = (execute M (start_config k zs) (t + fst (xs ! j)) <:> j) i" using exec_def by simp then show ?thesis using tape_alphabet[OF tm_M] zs_less_G by (simp add: "1" nth_append) next case 2 then have "?ys ! j = (case snd (xs ! (j-k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j-k) then 1 else 0)" by (simp add: less_2k_nth) then have "?ys ! j \ 1" by (cases "snd (xs ! (j - k))") auto then show ?thesis using G_ge_4 by simp next case 3 then have "?ys ! j \ 1" by (simp add: twok_nth) then show ?thesis using G_ge_4 by simp next case 4 then have "?ys ! j = (if i = 0 then 1 else 0)" using twok1_nth[of ?f1 ?f2 ?a ?b] by simp then show ?thesis using G_ge_4 by simp qed qed moreover have "length ?ys = 2 * k + 2" by simp ultimately show "\x\set ?ys. x < G" by (metis (no_types, lifting) in_set_conv_nth) qed lemma dec_zip_cont: assumes "i < TT" shows "dec (zip_cont t xs i) = (map (\j. (exec (t + fst (xs ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..x\set ?ys. x < G" using zip_cont_less_G[OF assms] by simp moreover have len: "length ?ys = 2 * k + 2" by simp ultimately have "dec (enc ?ys) = ?ys" using dec_enc by simp then show ?thesis using zip_cont_def assms by simp qed lemma zip_cont_gt_1: assumes "i < TT" shows "zip_cont t xs i > 1" using assms enc_greater zip_cont_def by simp lemma zip_cont_less: assumes "i < TT" shows "zip_cont t xs i < G ^ (2 * k + 2) + 2" using assms enc_less zip_cont_less_G zip_cont_def by simp lemma zip_cont_eq_0: assumes "i \ TT" shows "zip_cont t xs i = 0" using assms zip_cont_def by simp lemma dec_zip_cont_less_k: assumes "i < TT" and "j < k" shows "dec (zip_cont t xs i) ! j = (exec (t + fst (xs ! j)) <:> j) i" using dec_zip_cont[OF assms(1)] using assms(2) less_k_nth by (simp add: less_k_nth) lemma dec_zip_cont_less_2k: assumes "i < TT" and "j \ k" and "j < 2 * k" shows "dec (zip_cont t xs i) ! j = (case snd (xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using dec_zip_cont[OF assms(1)] assms(2,3) by (simp add: less_2k_nth) lemma dec_zip_cont_2k: assumes "i < TT" shows "dec (zip_cont t xs i) ! (2 * k) = (if i < t then 1 else 0)" using dec_zip_cont[OF assms(1)] by (simp add: twok_nth) lemma dec_zip_cont_2k1: assumes "i < TT" shows "dec (zip_cont t xs i) ! (2 * k + 1) = (if i = 0 then 1 else 0)" using dec_zip_cont[OF assms(1)] twok1_nth by force lemma zip_cont_eqI: assumes "i < TT" and "\j. j < k \ (exec (t + fst (xs ! j)) <:> j) i = (exec (t + fst (xs' ! j)) <:> j) i" and "\j. j < k \ (case snd (xs ! j) of None \ (0::nat) | Some d \ if i = exec (t + d) <#> j then 1 else 0) = (case snd (xs' ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0)" shows "zip_cont t xs i = zip_cont t xs' i" proof - have *: "map (\j. case snd (xs ! j) of None \ (0::nat) | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..j. case snd (xs' ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..j. (exec (t + fst (xs ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..j. (exec (t + fst (xs' ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..j. (exec (t + fst (xs' ! j)) <:> j) i) [0..j. case snd (xs' ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..j. j < k \ (exec (t + fst (xs ! j)) <:> j) i = (exec (t + fst (xs' ! j)) <:> j) i" and "\j. j < k \ snd (xs ! j) = snd (xs' ! j)" shows "zip_cont t xs i = zip_cont t xs' i" using assms zip_cont_eqI by presburger lemma begin_tape_zip_cont: "begin_tape {y. y < G ^ (2 * k + 2) + 2 \ 1 < y \ dec y ! (2 * k + 1) = 1} (zip_cont t xs, i)" (is "begin_tape ?ys _") proof - let ?y = "zip_cont t xs 0" have "?y \ ?ys" proof - have *: "?y = enc (map (\j. (exec (t + fst (xs ! j)) <:> j) 0) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if 0 = exec (t + d) <#> j then 1 else 0) [0.. 1" using enc_greater by simp have "?z ! (2 * k + 1) = 1" using twok1_nth by fast then have 2: "dec ?y ! (2 * k + 1) = 1" using dec_zip_cont by simp have "?y < G ^ (2 * k + 2) + 2" using enc_less * zip_cont_less_G[of 0] by simp then show ?thesis using 1 2 by simp qed moreover have "zip_cont t xs i \ ?ys" if "i > 0" for i proof (cases "i < TT") case True then have "dec (zip_cont t xs i) = (map (\j. (exec (t + fst (xs ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0.. es5 @ map (\i. (1 + Suc i, 1 + Suc i)) [0.. [(\zs\, n + 1), (zip_cont 0 (replicate k (0, Some 0)), n + 1)] @ replicate (2 * k + 1) (\\\)" lemma tm6: "traces tm6 tps0 es6 tps6" unfolding tm6_def es6_def proof (rule traces_sequential) show "traces tm5 tps0 es5 tps5" using tm5 . show "traces tm5_6 tps5 (map (\i. (1 + Suc i, 1 + Suc i)) [0..i. (1 + Suc i, 1 + Suc i)) [0..i. (tps5 :#: 0 + Suc i, tps5 :#: 1 + Suc i)) [0..h. enc (h mod G # replicate (k - 1) 0 @ replicate k 0 @ [0, 0])) n]" proof - define f where "f = (\h. enc (h mod G # replicate (k - 1) 0 @ replicate k 0 @ [0, 0]))" let ?tp1 = "tps5 ! 0" let ?tp2 = "tps5 ! 1" let ?xs = "replicate k (0::nat, Some 0::nat option)" have rhs_less_TT: "zip_cont 0 ?xs i = enc (map (\j. (start_config k zs <:> j) i) [0..j. if i = start_config k zs <#> j then 1 else 0) [0..j. (exec (0 + fst (?xs ! j)) <:> j) i) [0..j. case snd (?xs ! j) of None \ 0 | Some d \ if i = exec (0 + d) <#> j then 1 else 0) [0..j. (exec (0 + fst (?xs ! j)) <:> j) i) [0..j. (exec 0 <:> j) i) [0..j. (exec 0 <:> j) i) [0..j. case snd (?xs ! j) of None \ 0 | Some d \ if i = exec (0 + d) <#> j then 1 else 0) [0..j. (exec 0 <:> j) i) [0..j. if i = exec 0 <#> j then 1 else 0) [0..j. case snd (?xs ! j) of None \ 0 | Some d \ if i = exec (0 + d) <#> j then 1 else 0) [0..j. if i = exec 0 <#> j then 1 else 0) [0.. i \ i < snd ?tp2 + n then f (fst ?tp1 (snd ?tp1 + i - snd ?tp2)) else fst ?tp2 i) = zip_cont 0 (replicate k (0, Some 0)) i" (is "?lhs = ?rhs") for i proof (cases "1 \ i \ i < 1 + n") case True then have "snd ?tp2 \ i \ i < snd ?tp2 + n" using tps5_def by simp then have "?lhs = f (fst ?tp1 (snd ?tp1 + i - snd ?tp2))" by simp then have "?lhs = f (fst ?tp1 i)" using tps5_def by simp then have "?lhs = f (zs ! (i - 1))" (is "_ = f (zs ! ?i)") using tps5_def contents_def True by simp moreover have "zs ! ?i < G" using True zs_less_G by auto ultimately have lhs: "?lhs = enc (zs ! ?i # replicate (k - 1) 0 @ replicate k 0 @ [0, 0])" using f_def by simp have "TT > n" using fmt_ge_n by (simp add: le_imp_less_Suc) then have "i < TT" using True by simp then have rhs: "?rhs = enc (map (\j. (start_config k zs <:> j) i) [0..j. if i = start_config k zs <#> j then 1 else 0) [0..j. (start_config k zs <:> j) i) [0.. 0) i = zs ! ?i" using start_config_def True by simp then show ?thesis using c1 that by auto next case c2: False then have "(start_config k zs <:> j) i = 0" using start_config_def True that by simp then show ?thesis using c2 that by simp qed qed moreover have "map (\j. if i = start_config k zs <#> j then 1 else 0) [0.. j = 0" if "j < k" for j using that start_config_pos by auto then have "map (\j. if i = start_config k zs <#> j then 1 else 0) [0..j. 0) [0..j. (start_config k zs <:> j) i) [0..j. if i = start_config k zs <#> j then 1 else 0) [0..j. (start_config k zs <:> j) i) [0..j. if i = start_config k zs <#> j then 1 else 0) [0.. n" using outside by simp then have lhs: "?lhs = fst ?tp2 i" using tps5_def by simp then show ?thesis proof (cases "i < TT") case True moreover have "i > 0" using False by simp ultimately have lhs: "?lhs = enc (replicate k 0 @ replicate k 0 @ [0, 0])" using lhs tps5_def contents'_def by simp have "?rhs = enc (map (\j. (start_config k zs <:> j) i) [0..j. if i = start_config k zs <#> j then 1 else 0) [0..j. (start_config k zs <:> j) i) [0.. n` by simp next case False then show ?thesis using that start_config_def `i > 0` by simp qed qed moreover have "map (\j. if i = start_config k zs <#> j then 1 else 0) [0.. j = 0" if "j < k" for j using that start_config_pos by auto then have "map (\j. if i = start_config k zs <#> j then 1 else 0) [0..j. 0) [0.. TT" by simp moreover have "length (enc (replicate k 1 @ replicate k 1 @ [0, 1]) # replicate (fmt n) (enc (replicate k 0 @ replicate k 0 @ [0, 0]))) = TT" by simp ultimately have "?lhs = 0" using lhs contents'_at_ge by (simp add: tps5_def) moreover have "?rhs = 0" using zip_cont_def `i \ TT` by simp ultimately show ?thesis by simp qed qed qed then have "(\i. if snd ?tp2 \ i \ i < snd ?tp2 + n then f (fst ?tp1 (snd ?tp1 + i - snd ?tp2)) else fst ?tp2 i) = zip_cont 0 (replicate k (0, Some 0))" by simp moreover have "transplant (tps5 ! 0) (tps5 ! 1) (\h. enc (h mod G # replicate (k - 1) 0 @ replicate k 0 @ [0, 0])) n = (\i. if snd ?tp2 \ i \ i < snd ?tp2 + n then f (fst ?tp1 (snd ?tp1 + i - snd ?tp2)) else fst ?tp2 i, snd ?tp2 + n)" using transplant_def f_def by auto ultimately have "transplant (tps5 ! 0) (tps5 ! 1) (\h. enc (h mod G # replicate (k - 1) 0 @ replicate k 0 @ [0, 0])) n = (zip_cont 0 (replicate k (0, Some 0)), n + 1)" using tps5_def by simp then have "tps6 ! 1 = transplant (tps5 ! 0) (tps5 ! 1) (\h. enc (h mod G # replicate (k - 1) 0 @ replicate k 0 @ [0, 0])) n" using tps6_def by simp moreover have "tps6 ! 0 = tps5 ! 0 |+| n" using tps6_def tps5_def by simp ultimately show ?thesis using tps5_def tps6_def by simp qed qed qed definition tps7 :: "tape list" where "tps7 \ [(\zs\, n + 1), (zip_cont 0 (replicate k (0, Some 0)), 0)] @ replicate (2 * k + 1) (\\\)" definition "es7 \ es6 @ map (\i. (n + 1, i)) (rev [0.. [(\zs\, n + 1), (zip_cont 0 (replicate k (0, Some 0)), 0), \\\] @ replicate (2 * k) (\\\)" definition "es8 \ es7 @ [(n + 1, 0)]" lemma len_es8: "length es8 = length (es_fmt n) + 2 * TT + 2 * n + 7" using es8_def len_es7 by simp lemma tm8: "traces tm8 tps0 es8 tps8" unfolding tm8_def es8_def proof (rule traces_sequential) show "traces tm7 tps0 es7 tps7" using tm7 . show "traces (tm_write 2 0) tps7 [(n + 1, 0)] tps8" proof (rule traces_tm_write_ge2I) show "(2::nat) \ 2" by simp show "2 < length tps7" "[(n + 1, 0)] = [(tps7 :#: 0, tps7 :#: 1)]" using tps7_def by simp_all show "tps8 = tps7[2 := tps7 ! 2 |:=| 0]" proof (rule nth_equalityI) show "length tps8 = length (tps7[2 := tps7 ! 2 |:=| 0])" using tps7_def tps8_def by simp show "tps8 ! i = tps7[2 := tps7 ! 2 |:=| 0] ! i" if "i < length tps8" for i proof - consider "i = 0" | "i = 1" | "i = 2" | "i > 2" by linarith then show ?thesis proof (cases) case 1 then show ?thesis using tps7_def tps8_def by simp next case 2 then show ?thesis using tps7_def tps8_def by simp next case 3 then have *: "tps8 ! i = \\\" using tps8_def by simp have "(tps7 ! 2) |:=| \ = \\\" using tps7_def onesie_write by simp then have "(tps7[2 := tps7 ! 2 |:=| \]) ! 2 = \\\" using tps7_def by simp then show ?thesis using 3 * by simp next case 4 then have "tps8 ! i = \\\" using tps8_def that by simp moreover have "tps7 ! i = \\\" using tps7_def that "4" tps8_def by auto ultimately show ?thesis by (simp add: "4" less_not_refl3) qed qed qed qed qed definition tps9 :: "tape list" where "tps9 \ [(\zs\, n + 1), (zip_cont 0 (replicate k (0, Some 0)), 0), \\\] @ replicate (2 * k) (\\\)" definition "es9 \ es8 @ [(n + 1, 0)]" lemma len_es9: "length es9 = length (es_fmt n) + 2 * TT + 2 * n + 8" using es9_def len_es8 by simp lemma tm9: "traces tm9 tps0 es9 tps9" unfolding tm9_def es9_def proof (rule traces_sequential[OF tm8]) show "traces (tm_write_many {3..<2 * k + 3} 0) tps8 [(n + 1, 0)] tps9" proof (rule traces_tm_write_manyI[where ?k="2*k+3"]) show "0 \ {3..<2 * k + 3}" by simp show "\j\ {3..<2 * k + 3}. j < 2 * k + 3" by simp show "2 \ 2 * k + 3" by simp show "length tps8 = 2 * k + 3" "length tps9 = 2 * k + 3" using tps8_def tps9_def by simp_all show "[(n + 1, 0)] = [(tps8 :#: 0, tps8 :#: 1)]" using tps8_def by simp show "tps9 ! j = tps8 ! j" if "j < 2 * k + 3" "j \ {3..<2 * k + 3}" for j proof - have "j < 3" using that by simp then show ?thesis using tps8_def tps9_def by (metis (no_types, lifting) length_Cons list.size(3) nth_append numeral_3_eq_3) qed show "tps9 ! j = tps8 ! j |:=| 0" if "j \ {3..<2 * k + 3}" for j proof - have j: "j \ 3" "j < 2 * k + 3" using that by simp_all then have "tps8 ! j = \\\" using tps8_def by simp moreover have "tps9 ! j = \\\" using j tps9_def by simp ultimately show ?thesis by (simp add: onesie_write) qed qed qed subsubsection \The loop\ text \ Immediately before and during the loop the tapes will have the shape below. The input tape will stay unchanged. The output tape will contain the $k$ encoded tapes of $M$. The first memorization tape will contain $M$'s state. The following $k$ memorization tapes will store information about head movements. The final $k$ memorization tapes will have information about read or to-be-written symbols. \ definition tpsL :: "nat \ (nat \ nat option) list \ nat \ nat \ (nat \ nat) \ (nat \ symbol) \ tape list" where "tpsL t xs i q mvs symbs \ (\zs\, n + 1) # (zip_cont t xs, i) # \fst (exec (t + q))\ # map (onesie \ mvs) [0.. symbs) [0.." unfolding tpsL_def using contents_def read_def by simp lemma read_tpsL_1: "read (tpsL t xs i q mvs symbs) ! 1 = (if i < TT then enc (map (\j. (exec (t + fst (xs ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0.. j) i" using assms read_tpsL_1 dec_zip_cont_less_k enc_nth_def zip_cont_def by auto lemma read_tpsL_1_nth_less_2k: assumes "i < TT" and "k \ j" and "j < 2 * k" shows "enc_nth (read (tpsL t xs i q mvs symbs) ! 1) j = (case snd (xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using assms read_tpsL_1 dec_zip_cont_less_2k enc_nth_def zip_cont_def by simp lemma read_tpsL_1_nth_2k: assumes "i < TT" shows "enc_nth (read (tpsL t xs i q mvs symbs) ! 1) (2 * k) = (if i < t then 1 else 0)" using assms read_tpsL_1 dec_zip_cont_2k enc_nth_def zip_cont_def by simp lemma read_tpsL_1_nth_2k1: assumes "i < TT" shows "enc_nth (read (tpsL t xs i q mvs symbs) ! 1) (2 * k + 1) = (if i = 0 then 1 else 0)" using assms read_tpsL_1 dec_zip_cont_2k1 enc_nth_def zip_cont_def by simp lemma read_tpsL_1_bounds: assumes "i < TT" shows "read (tpsL t xs i q mvs symbs) ! 1 > 1" and "read (tpsL t xs i q mvs symbs) ! 1 < G ^ (2 * k + 2) + 2" proof - have "read (tpsL t xs i q mvs symbs) ! 1 = zip_cont t xs i" using tpsL_def read_def by simp then show "read (tpsL t xs i q mvs symbs) ! 1 > 1" and "read (tpsL t xs i q mvs symbs) ! 1 < G ^ (2 * k + 2) + 2" using zip_cont_gt_1[OF assms] zip_cont_less[OF assms] by simp_all qed lemma read_tpsL_2: "read (tpsL t xs i q mvs symbs) ! 2 = fst (exec (t + q))" unfolding tpsL_def using contents_def read_def by simp lemma read_tpsL_3: assumes "3 \ j" and "j < k + 3" shows "read (tpsL t xs i q mvs symbs) ! j = mvs (j - 3)" proof - define j' where "j' = j - 3" then have "j' < k" using assms by simp have "read (tpsL t xs i q mvs symbs) ! j = (map (tape_read \ (onesie \ mvs)) [0.. (onesie \ symbs)) [0.. (onesie \ mvs)) [0.. (onesie \ symbs)) [0.. j" and "j < 2 * k + 3" shows "read (tpsL t xs i q mvs symbs) ! j = symbs (j - k - 3)" proof - define j' where "j' = j - 3" then have j': "k \ j'" "j' < k + k" using assms by simp_all have "read (tpsL t xs i q mvs symbs) ! j = (map (tape_read \ (onesie \ mvs)) [0.. (onesie \ symbs)) [0.. (onesie \ mvs)) [0.. (onesie \ symbs)) [0.. (\_. c)) [0..c\" by (metis List.map.compositionality map_replicate map_replicate_trivial) text \ After the initialization, that is, right before the loop, the simulator's tapes look like this: \ lemma tps9_tpsL: "tps9 = tpsL 0 (replicate k (0, Some 0)) 0 0 (\j. 0) (\j. 0)" proof - have "fst (exec 0) = 0" using exec_def by (simp add: start_config_def) then have "tpsL 0 (replicate k (0, Some 0)) 0 0 (\j. 0) (\j. 0) = (\zs\, n + 1) # (zip_cont 0 (replicate k (0, Some 0)), 0) # \\\ # replicate k (\\\) @ replicate k (\\\)" using tpsL_def map_const_upt by simp then show ?thesis using tps9_def by (metis Cons_eq_appendI mult_2 replicate_add self_append_conv2) qed lemma tpsL_0: "tpsL t xs i q mvs symbs ! 0 = (\zs\, n + 1)" using tpsL_def by simp lemma tpsL_1: "tpsL t xs i q mvs symbs ! 1 = (zip_cont t xs, i)" using tpsL_def by simp lemma tpsL_2: "tpsL t xs i q mvs symbs ! 2 = \fst (exec (t + q))\" using tpsL_def by simp lemma tpsL_mvs: "j < k \ tpsL t xs i q mvs symbs ! (3 + j) = \mvs j\" using tpsL_def by (simp add: nth_append) lemma tpsL_mvs': "3 \ j \ j < 3 + k \ tpsL t xs i q mvs symbs ! j = \mvs (j - 3)\" using tpsL_mvs by (metis add.commute le_add_diff_inverse less_diff_conv2) lemma tpsL_symbs: assumes "j < k" shows "tpsL t xs i q mvs symbs ! (3 + k + j) = \symbs j\" using tpsL_def assms by (simp add: nth_append) lemma tpsL_symbs': assumes "3 + k \ j" and "j < 2 * k + 3" shows "tpsL t xs i q mvs symbs ! j = \symbs (j - k - 3)\" proof - have "j - (k + 3) < k" using assms(1) assms(2) by linarith then show ?thesis using assms(1) tpsL_symbs by fastforce qed text \The condition: less than $TT$ steps simulated.\ definition tpsC0 :: "nat \ tape list" where "tpsC0 t \ tpsL t (replicate k (0, Some 0)) 0 0 (\j. 0) (\j. 0)" definition "tpsC1 t \ tpsL t (replicate k (0, Some 0)) t 0 (\j. 0) (\j. 0)" definition "esC t \ map (\i. (n + 1, Suc i)) [0.. set (filter f [0.. x < N" by simp lemma set_filter_upt': "x \ set (filter f [0.. f x" by simp text \ We will use this TM at the end of the loop again. Therefore we prove a more general result than necessary at this point. \ lemma tmC_general: assumes "t \ TT" and "tps = tpsL t xs 0 i mvs symbs" and "tps' = tpsL t xs t i mvs symbs" shows "traces tmC tps (esC t) tps'" unfolding tmC_def proof (rule traces_tm_right_until_1I[where ?n="t"]) show "1 < length tps" using assms(2) by simp show "tps' = tps[1 := tps ! 1 |+| t]" using assms(2,3) tpsL_def by simp show "esC t = map (\i. (tps :#: 0, tps :#: 1 + Suc i)) [0.. (y = 0 \ dec y ! (2 * k) = 0)} t" (is "rneigh _ ?s t") proof (rule rneighI) have 1: "tps ! 1 = (zip_cont t xs, 0)" using assms(2) tpsL_def by simp have s: "?s = {y. y = 0 \ (dec y ! (2 * k) = 0 \ y < G ^ (2 * k + 2) + 2)}" (is "_ = ?r") by auto show "(tps ::: 1) (tps :#: 1 + t) \ ?s" proof (cases "t = TT") case True then have "tps :#: 1 + t = TT" using assms(2) tpsL_def by simp moreover have "(tps ::: 1) TT = 0" using assms(2) tpsL_def zip_cont_def by simp ultimately show ?thesis by auto next case False with assms have "t < TT" by simp let ?y = "(tps ::: 1) t" have "dec ?y ! (2 * k) = 0" using assms(2) tpsL_def dec_zip_cont[OF `t < TT`] by (simp add: twok_nth) moreover have "?y < G ^ (2 * k + 2) + 2" using assms(2) tpsL_def `t < TT` zip_cont_less by simp ultimately have "?y \ ?s" using s by simp then show ?thesis using 1 by simp qed show "(tps ::: 1) (tps :#: 1 + m) \ ?s" (is "?y \ ?s") if "m < t" for m proof - have "m < TT" using that assms by simp then have "dec ?y ! (2 * k) = 1" using tpsC0_def tpsL_def dec_zip_cont[OF `m < TT`] that by (metis (no_types, lifting) "1" add_cancel_right_left dec_zip_cont_2k prod.sel(1) prod.sel(2)) moreover from `m < TT` have "?y > 1" using 1 zip_cont_gt_1 by simp ultimately show ?thesis using s by simp qed qed qed corollary tmC: assumes "t \ TT" shows "traces tmC (tpsC0 t) (esC t) (tpsC1 t)" using tmC_general tpsC0_def tpsC1_def assms by simp lemma tpsC1_at_T: "tpsC1 TT :.: 1 = 0" using tpsC1_def tpsL_def zip_cont_def by simp lemma tpsC1_at_less_T: "t < TT \ tpsC1 t :.: 1 > 0" proof - assume "t < TT" then have "tpsC1 t :.: 1 > 1" using tpsC1_def tpsL_def zip_cont_def enc_greater by simp then show ?thesis by simp qed text \The body of the loop: simulating one step\ definition "tpsL0 t \ tpsL t (replicate k (0, Some 0)) t 0 (\j. 0) (\j. 0)" lemma tpsL0_eq_tpsC1: "tpsL0 t = tpsC1 t" using tpsL0_def tpsC1_def by simp definition "esL1 t \ map (\i. (n + 1, i)) (rev [0.. tpsL t (replicate k (0, Some 0)) 0 0 (\j. 0) (\j. 0)" lemma tmL1: "traces tmL1 (tpsL0 t) (esL1 t) (tpsL1 t)" unfolding tmL1_def esL1_def proof (rule traces_tm_left_until_1I) show "1 < length (tpsL0 t)" using tpsL0_def tpsL_def by simp show "map (Pair (n + 1)) (rev [0..Collecting the read symbols of the simulated machines:\ lemma sem_cmdL2_ge_TT: assumes "tps = tpsL t xs i q mvs symbs" and "i \ TT" shows "sem cmdL2 (0, tps) = (1, tps)" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) cmdL2" using cmdL2_def by simp show len: "length tps = 2 * k + 3" using assms(1) by simp show "length tps = 2 * k + 3" using assms(1) by simp let ?rs = "read tps" show "fst (cmdL2 ?rs) = 1" proof - have "?rs ! 1 = \" using assms read_tpsL_1 by auto then show ?thesis by (simp add: cmdL2_def) qed then show "act (cmdL2 ?rs [!] i) (tps ! i) = tps ! i" if "i < 2 * k + 3" for i using assms that by (metis (no_types, lifting) One_nat_def act_Stay cmdL2_at_eq_0 cmdL2_def len less_Suc_eq_0_disj less_numeral_extra(4) prod.sel(1) read_length) qed lemma sem_cmdL2_less_TT: assumes "tps = tpsL t xs i q mvs symbs" and "symbs = (\j. if exec t <#> j < i then exec t <.> j else 0) " and "tps' = tpsL t xs (Suc i) q mvs symbs'" and "symbs' = (\j. if exec t <#> j < Suc i then exec t <.> j else 0) " and "i < TT" and "xs = replicate k (0, Some 0)" shows "sem cmdL2 (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) cmdL2" using cmdL2_def by simp show len: "length tps = 2 * k + 3" using assms(1) by simp show len': "length tps' = 2 * k + 3" using assms(3) by simp define rs where "rs = read tps" then have rs_at_1: "rs ! 1 \ \" using assms(1,5) read_tpsL_1 enc_greater by (metis (no_types, lifting) not_one_less_zero) then show "fst (cmdL2 (read tps)) = 0" by (simp add: cmdL2_def rs_def) show "act (cmdL2 (read tps) [!] j) (tps ! j) = tps' ! j" if "j < 2 * k + 3" for j proof - have snd: "snd (cmdL2 rs) = [(rs!0, Stay), (rs!1, Right), (rs!2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (if 1 < rs ! 1 \ rs ! 1 < G^(2*k+2)+2 \ enc_nth (rs!1) (k+j) = 1 then enc_nth (rs!1) j else rs!(3+k+j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using \j < 2 * k + 3\ by linarith then show ?thesis proof cases case 1 then have "cmdL2 rs [!] j = (rs ! 0, Stay)" by (simp add: snd) then show ?thesis using act_Stay assms(1,3) tpsL_def "1" rs_def read_tpsL_0 by auto next case 2 then have "cmdL2 rs [!] j = (rs ! 1, Right)" by (simp add: snd) then show ?thesis using act_Right assms(1,3) "2" rs_def by (metis Suc_eq_plus1 add.commute add_Suc fst_conv len less_add_Suc1 numeral_3_eq_3 sndI tpsL_1) next case 3 then have "cmdL2 rs [!] j = (rs ! 2, Stay)" by (simp add: snd) then show ?thesis using act_Stay assms(1,3) "3" rs_def by (metis len that tpsL_2) next case 4 then have "cmdL2 rs [!] j = (rs ! j, Stay)" using cmdL2_at_3 rs_at_1 by simp then show ?thesis using act_Stay assms(1,3) "4" rs_def by (metis add.commute len that tpsL_mvs') next case 5 then have 1: "cmdL2 rs [!] j = (if 1 < rs ! 1 \ rs ! 1 < G^(2*k+2)+2 \ enc_nth (rs ! 1) (j - 3) = 1 then enc_nth (rs ! 1) (j - k - 3) else rs ! j, Stay)" using cmdL2_at_4 rs_at_1 by simp have enc: "rs ! 1 = enc (map (\j. (exec (t + fst (xs ! j)) <:> j) i) [0..j. case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0.. 1" "rs ! 1 < G ^ (2 * k + 2) + 2" using rs_def assms(1,5) read_tpsL_1_bounds by simp_all then have cmd1: "cmdL2 rs [!] j = (if enc_nth (rs ! 1) (j - 3) = 1 then enc_nth (rs ! 1) (j - k - 3) else rs ! j, Stay)" using 1 by simp have "enc_nth (rs ! 1) (j - 3) = (case snd (xs ! (j - 3 - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - 3 - k) then 1 else 0)" using read_tpsL_1_nth_less_2k[where ?j="j - 3"] 5 assms(1,5) rs_def by auto then have "enc_nth (rs ! 1) (j - 3) = (if i = exec t <#> (j - 3 - k) then 1 else 0)" using 5 assms(6) by auto then have cmd2: "enc_nth (rs ! 1) (j - 3) = (if i = exec t <#> (j - k - 3) then 1 else 0)" by (metis diff_right_commute) let ?j = "j - k - 3" have "enc_nth (rs ! 1) ?j = (exec (t + fst (xs ! ?j)) <:> ?j) i" using read_tpsL_1_nth_less_k[where ?j="j - k - 3"] 5 assms(1,5) rs_def by auto then have "enc_nth (rs ! 1) ?j = (exec t <:> ?j) i" using assms(6) 5 by auto then have "cmdL2 rs [!] j = (if i = exec t <#> ?j then (exec t <:> ?j) i else rs ! j, Stay)" using cmd1 cmd2 by simp then have command: "cmdL2 rs [!] j = (if i = exec t <#> ?j then exec t <.> ?j else rs ! j, Stay)" by simp have tps: "tps ! j = \if exec t <#> ?j < i then exec t <.> ?j else \\" using 5 assms(1,2) tpsL_symbs' by simp have tps': "tps' ! j = \if exec t <#> ?j < Suc i then exec t <.> ?j else \\" using 5 assms(3,4) tpsL_symbs' by simp show "act (cmdL2 (read tps) [!] j) (tps ! j) = tps' ! j" proof (cases "exec t <#> ?j = i") case True then have "act (cmdL2 (read tps) [!] j) (tps ! j) = act (exec t <.> ?j, Stay) (tps ! j)" using rs_def command by simp also have "... = act (exec t <.> ?j, Stay) \if exec t <#> ?j < i then exec t <.> ?j else \\" using tps by simp also have "... = \exec t <.> ?j\" using act_onesie by simp also have "... = \if exec t <#> ?j < Suc i then exec t <.> ?j else \\" using True by simp also have "... = tps' ! j" using tps' by simp finally show ?thesis . next case False then have "act (cmdL2 (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using rs_def command by simp also have "... = tps ! j" using rs_def act_Stay by (simp add: "5" len) also have "... = \if exec t <#> ?j < i then exec t <.> ?j else \\" using tps by simp also have "... = \if exec t <#> ?j < Suc i then exec t <.> ?j else \\" using False by simp also have "... = tps' ! j" using tps' by simp finally show ?thesis . qed qed qed qed corollary sem_cmdL2_less_Tfmt: assumes "xs = replicate k (0, Some 0)" and "i < TT" shows "sem cmdL2 (0, tpsL t xs i q mvs (\j. if exec t <#> j < i then exec t <.> j else \)) = (0, tpsL t xs (Suc i) q mvs (\j. if exec t <#> j < Suc i then exec t <.> j else \))" using sem_cmdL2_less_TT assms by simp lemma execute_cmdL2_le_TT: assumes "tt \ TT" and "xs = replicate k (0, Some 0)" and "tps = tpsL t xs 0 q mvs (\_. \)" shows "execute tmL1_2 (0, tps) tt = (0, tpsL t xs tt q mvs (\j. if exec t <#> j < tt then exec t <.> j else \))" using assms(1) proof (induction tt) case 0 then show ?case using assms(3) by simp next case (Suc tt) then have "execute tmL1_2 (0, tps) (Suc tt) = exe tmL1_2 (execute tmL1_2 (0, tps) tt)" by simp also have "... = exe tmL1_2 (0, tpsL t xs tt q mvs (\j. if exec t <#> j < tt then exec t <.> j else \))" using Suc by simp also have "... = sem cmdL2 (0, tpsL t xs tt q mvs (\j. if exec t <#> j < tt then exec t <.> j else \))" unfolding tmL1_2_def using Suc by (simp add: exe_lt_length) also have "... = (0, tpsL t xs (Suc tt) q mvs (\j. if exec t <#> j < Suc tt then exec t <.> j else \))" using sem_cmdL2_less_Tfmt[OF assms(2)] Suc by simp finally show ?case . qed lemma tpsL_symbs_eq: assumes "\j. j < k \ symbs j = symbs' j" shows "tpsL t xs i q mvs symbs = tpsL t xs i q mvs symbs'" unfolding tpsL_def using assms by simp lemma execute_cmdL2_Suc_TT: assumes "xs = replicate k (0, Some 0)" and "tps = tpsL t xs 0 q mvs (\j. 0)" and "t < TT" shows "execute tmL1_2 (0, tps) (Suc TT) = (1, tpsL t xs TT q mvs (\j. exec t <.> j))" proof - have "execute tmL1_2 (0, tps) (Suc TT) = exe tmL1_2 (execute tmL1_2 (0, tps) TT)" by simp also have "... = exe tmL1_2 (0, tpsL t xs TT q mvs (\j. if exec t <#> j < TT then exec t <.> j else \))" using execute_cmdL2_le_TT[where ?tt=TT] assms(1,2) by simp also have "... = sem cmdL2 (0, tpsL t xs TT q mvs (\j. if exec t <#> j < TT then exec t <.> j else \))" unfolding tmL1_2_def by (simp add: exe_lt_length) also have "... = (1, tpsL t xs TT q mvs (\j. if exec t <#> j < TT then exec t <.> j else \))" using sem_cmdL2_ge_TT[where ?i=TT] by simp finally have "execute tmL1_2 (0, tps) (Suc TT) = (1, tpsL t xs TT q mvs (\j. if exec t <#> j < TT then exec t <.> j else \))" . moreover have "(\j. if exec t <#> j < TT then exec t <.> j else \) j = (\j. exec t <.> j) j" if "j < k" for j using exec_pos_less_TT assms(3) that by simp ultimately show ?thesis using tpsL_symbs_eq by fastforce qed definition "esL1_2 \ map (\i. (n + 1, Suc i)) [0.._. \)) esL1_2 (tpsL t xs TT q mvs (\j. exec t <.> j))" proof have len: "length esL1_2 = Suc TT" unfolding esL1_2_def by simp let ?tps = "tpsL t xs 0 q mvs (\j. 0)" show "execute tmL1_2 (0, ?tps) (length esL1_2) = (length tmL1_2, tpsL t xs (Suc (fmt n)) q mvs (\j. exec t <.> j))" using len execute_cmdL2_Suc_TT[OF assms(1) _ assms(2)] by (simp add: tmL1_2_def) show "\i. i < length esL1_2 \ fst (execute tmL1_2 (0, ?tps) i) < length tmL1_2" using len tmL1_2_def execute_cmdL2_le_TT assms(1) by (metis (no_types, lifting) One_nat_def Pair_inject length_Cons less_Suc_eq_le less_one list.size(3) prod.collapse) show "snd (execute tmL1_2 (0, ?tps) (Suc i)) :#: 0 = fst (esL1_2 ! i) \ snd (execute tmL1_2 (0, ?tps) (Suc i)) :#: 1 = snd (esL1_2 ! i)" if "i < length esL1_2" for i proof (cases "i < TT") case True then have "execute tmL1_2 (0, ?tps) (Suc i) = (0, tpsL t xs (Suc i) q mvs (\j. if exec t <#> j < Suc i then exec t <.> j else \))" using execute_cmdL2_le_TT[of "Suc i" xs] assms by simp then have "snd (execute tmL1_2 (0, ?tps) (Suc i)) = tpsL t xs (Suc i) q mvs (\j. if exec t <#> j < Suc i then exec t <.> j else \)" by simp moreover have "esL1_2 ! i = (n + 1, Suc i)" unfolding esL1_2_def using True nth_append by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 add.commute add_Suc diff_zero length_map length_upt nth_map_upt) ultimately show ?thesis using tpsL_pos_0 tpsL_pos_1 by simp next case False then have "i = TT" using len that by simp then have "execute tmL1_2 (0, ?tps) (Suc i) = (1, tpsL t xs TT q mvs (\j. exec t <.> j))" using execute_cmdL2_Suc_TT assms by simp moreover have "esL1_2 ! i = (n + 1, TT)" using `i = TT` esL1_2_def by (metis (no_types, lifting) diff_zero length_map length_upt nth_append_length) ultimately show ?thesis using tpsL_pos_0 tpsL_pos_1 by auto qed qed definition "esL2 t \ esL1 t @ esL1_2" definition "tpsL2 t \ tpsL t (replicate k (0, Some 0)) TT 0 (\_. \) (\j. exec t <.> j)" lemma tmL2: assumes "t < TT" shows "traces tmL2 (tpsL0 t) (esL2 t) (tpsL2 t)" unfolding tmL2_def esL2_def proof (rule traces_sequential[OF tmL1]) show "traces tmL1_2 (tpsL1 t) esL1_2 (tpsL2 t)" using traces_tmL1_2[OF _ assms] by (simp add: tpsL1_def tpsL2_def) qed definition "sim_nextstate t \ (if fst (exec t) < length M then fst ((M ! (fst (exec t))) (config_read (exec t))) else fst (exec t))" lemma sim_nextstate: "fst (exec (Suc t)) = sim_nextstate t" proof (cases "fst (exec t) < length M") case True let ?cfg = "exec t" let ?rs = "config_read ?cfg" let ?cmd = "M ! (fst ?cfg)" have "exec (Suc t) = sem ?cmd ?cfg" using exec_def True by (simp add: exe_lt_length) then have 2: "fst (exec (Suc t)) = fst (sem ?cmd ?cfg)" by simp also have "... = fst (?cmd ?rs)" using sem' by simp also have "... = sim_nextstate t" using sim_nextstate_def True by simp finally show ?thesis . next case False then show ?thesis using exec_def exe_def sim_nextstate_def by simp qed definition "sim_write t \ (if fst (exec t) < length M then map fst (snd ((M ! (fst (exec t))) (config_read (exec t)))) else config_read (exec t))" lemma sim_write_nth: assumes "fst (exec t) < length M" and "j < k" shows "sim_write t ! j = ((M ! (fst (exec t))) (config_read (exec t)) [.] j)" proof - have "length (snd ((M ! (fst (exec t))) (config_read (exec t)))) = k" using assms turing_commandD(1) exec_def execute_num_tapes read_length start_config_length tm_M turing_machine_def by (metis add_gr_0 less_imp_add_positive nth_mem) then show ?thesis using sim_write_def assms by simp qed lemma sim_write_nth_else: assumes "\ (fst (exec t) < length M)" shows "sim_write t ! j = config_read (exec t) ! j" by (simp add: assms sim_write_def) lemma sim_write_nth_less_G: assumes "j < k" shows "sim_write t ! j < G" proof (cases "fst (exec t) < length M") case True then have *: "sim_write t ! j = (M ! (fst (exec t))) (config_read (exec t)) [.] j" using sim_write_nth assms by simp have "turing_command k (length M) G (M ! (fst (exec t)))" using tm_M True turing_machineD(3) by simp moreover have "\i j = (exec t <:> j)(exec t <#> j := sim_write t ! j)" proof (cases "fst (exec t) < length M") case True let ?cfg = "exec t" let ?rs = "config_read ?cfg" let ?cmd = "M ! (fst ?cfg)" have len_rs: "length ?rs = k" using assms exec_def execute_num_tapes start_config_length read_length tm_M by simp have "turing_command k (length M) G ?cmd" using True tm_M turing_machineD(3) by simp then have len: "length (snd (?cmd ?rs)) = k" using len_rs turing_commandD(1) by simp have "sim_write t = map fst (snd (?cmd ?rs))" using sim_write_def True by simp then have 1: "sim_write t ! j = ?cmd ?rs [.] j" by (simp add: assms len) have 2: "exec (Suc t) = sem ?cmd ?cfg" using exec_def True by (simp add: exe_lt_length) have "snd (sem ?cmd ?cfg) = map (\(a, tp). act a tp) (zip (snd (?cmd ?rs)) (snd ?cfg))" using sem' by simp then have "snd (sem ?cmd ?cfg) ! j = (\(a, tp). act a tp) ((snd (?cmd ?rs) ! j), (snd ?cfg ! j))" using len assms len_rs read_length by simp then have "sem ?cmd ?cfg j = act (snd (?cmd ?rs) ! j) (?cfg j)" by simp then have "sem ?cmd ?cfg <:> j = fst (act (snd (?cmd ?rs) ! j) (?cfg j))" by simp then have "sem ?cmd ?cfg <:> j = (?cfg <:> j)(?cfg <#> j := fst (snd (?cmd ?rs) ! j))" using act by simp then have "sem ?cmd ?cfg <:> j = (?cfg <:> j)(?cfg <#> j := ?cmd ?rs [.] j)" . then show ?thesis using 1 2 by simp next case False let ?cfg = "exec t" let ?rs = "config_read ?cfg" have len_rs: "length ?rs = k" using assms exec_def execute_num_tapes start_config_length read_length tm_M by simp then have 1: "sim_write t ! j = ?rs ! j" using False by (simp add: sim_write_def) have 2: "?rs ! j = exec t <.> j" using assms len_rs read_abbrev read_length by auto have "exec (Suc t) = exec t" using exec_def exe_def False by simp then have "exec (Suc t) <:> j = exec t <:> j" by simp then show ?thesis using 1 2 by simp qed corollary sim_write': assumes "j < k" shows "(exec (Suc t) <:> j) (exec t <#> j) = sim_write t ! j" using assms sim_write by simp definition "sim_move t \ map direction_to_symbol (if fst (exec t) < length M then map snd (snd ((M ! (fst (exec t))) (config_read (exec t)))) else replicate k Stay)" lemma sim_move_nth: assumes "fst (exec t) < length M" and "j < k" shows "sim_move t ! j = direction_to_symbol ((M ! (fst (exec t))) (config_read (exec t)) [~] j)" proof - have "k = ||execute M (start_config k zs) t||" by (metis (no_types) Suc_1 execute_num_tapes start_config_length less_le_trans tm_M turing_machine_def zero_less_Suc) then show ?thesis by (smt (verit, best) turing_commandD(1) assms(1,2) exec_def length_map nth_map read_length sim_move_def tm_M turing_machineD(3)) qed lemma sim_move_nth_else: assumes "\ (fst (exec t) < length M)" and "j < k" shows "sim_move t ! j = 1" using assms sim_move_def direction_to_symbol_def by simp lemma sim_move: assumes "j < k" shows "exec (Suc t) <#> j = exec t <#> j + sim_move t ! j - 1" proof (cases "fst (exec t) < length M") case True let ?cfg = "exec t" let ?rs = "config_read ?cfg" let ?cmd = "M ! (fst ?cfg)" have len_rs: "length ?rs = k" using assms exec_def execute_num_tapes start_config_length read_length tm_M by simp have "turing_command k (length M) G ?cmd" using True tm_M turing_machineD(3) by simp then have len: "length (snd (?cmd ?rs)) = k" using len_rs turing_commandD(1) by simp have 1: "sim_move t ! j = direction_to_symbol (?cmd ?rs [~] j)" using assms sim_move_nth True by simp have "exec (Suc t) = sem ?cmd ?cfg" using exec_def True by (simp add: exe_lt_length) then have 2: "exec (Suc t) <#> j = sem ?cmd ?cfg <#> j" by simp have "snd (sem ?cmd ?cfg) = map (\(a, tp). act a tp) (zip (snd (?cmd ?rs)) (snd ?cfg))" using sem' by simp then have "snd (sem ?cmd ?cfg) ! j = (\(a, tp). act a tp) ((snd (?cmd ?rs) ! j), (snd ?cfg ! j))" using len assms len_rs read_length by simp then have "sem ?cmd ?cfg j = act (snd (?cmd ?rs) ! j) (?cfg j)" by simp then have "sem ?cmd ?cfg <#> j = snd (act (snd (?cmd ?rs) ! j) (?cfg j))" by simp then have "sem ?cmd ?cfg <#> j = (case ?cmd ?rs [~] j of Left \ ?cfg <#> j - 1 | Stay \ ?cfg <#> j | Right \ ?cfg <#> j + 1)" using act by simp then show ?thesis using 1 2 direction_to_symbol_def by (cases "?cmd ?rs [~] j") simp_all next case False then have "exec (Suc t) <#> j = exec t <#> j" using exec_def exe_def by simp moreover have "sim_move t ! j = 1" using direction_to_symbol_def sim_move_def assms False by simp ultimately show ?thesis by simp qed definition "tpsL3 t \ tpsL t (replicate k (0, Some 0)) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" lemma read_execute: "config_read (exec t) = map (\j. (exec t) <.> j) [0..j. exec t <.> j) [0.. i" unfolding rs_def tpsL2_def using read_tpsL_4[of "i + k + 3"] `i < k` by simp finally show ?thesis using `i < k` by simp qed qed then have drop: "drop (k + 3) rs = config_read (exec t)" using read_execute by simp then have drop_less_G: "\h\set (drop (k + 3) rs). h < G" using exec_def tm_M read_alphabet_set zs_less_G by presburger consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using \j < 2 * k + 3\ by linarith then show ?thesis proof (cases) case 1 then have "cmdL3 rs [!] j = (rs ! 0, Stay)" using cmdL3_def by simp moreover have "tpsL2 t ! j = (\zs\, n + 1)" using tpsL2_def 1 by (simp add: tpsL_0) moreover have "tpsL3 t ! j = (\zs\, n + 1)" using tpsL3_def 1 by (simp add: tpsL_0) ultimately show ?thesis using act_Stay by (metis 1 len rs_def that) next case 2 then have "cmdL3 rs [!] j = (rs ! 1, Stay)" using cmdL3_def by simp moreover have "tpsL2 t ! j = (zip_cont t (replicate k (0, Some 0)), TT)" using tpsL2_def 2 tpsL_1 by simp moreover have "tpsL3 t ! j = (zip_cont t (replicate k (0, Some 0)), TT)" using tpsL3_def 2 tpsL_1 by simp ultimately show ?thesis using act_Stay by (metis 2 len rs_def that) next case 3 show ?thesis proof (cases "rs ! 2 < length M") case True then have "cmdL3 rs [!] j = (fst ((M ! (rs ! 2)) (drop (k + 3) rs)), Stay)" using 3 drop_less_G cmdL3_at_2a by simp also have "... = (fst ((M ! (rs ! 2)) (config_read (exec t))), Stay)" using drop by simp also have "... = (fst (exec (Suc t)), Stay)" using True rs2 sim_nextstate sim_nextstate_def by auto finally have "cmdL3 rs [!] j = (fst (exec (Suc t)), Stay)" . then show ?thesis using tpsL2_def tpsL3_def tpsL_def 3 act_onesie rs_def by simp next case False then have "cmdL3 rs [!] j = (rs ! 2, Stay)" using 3 cmdL3_at_2b by simp moreover have "fst (exec (Suc t)) = rs ! 2" using False rs2 sim_nextstate sim_nextstate_def by auto ultimately have "cmdL3 rs [!] j = (fst (exec (Suc t)), Stay)" by simp then show ?thesis using tpsL2_def tpsL3_def tpsL_def 3 act_onesie rs_def by simp qed next case 4 then have 1: "j - 3 < k" by auto have 2: "tpsL2 t ! j = \\\" using tpsL2_def tpsL_mvs' 4 by simp have 3: "tpsL3 t ! j = \sim_move t ! (j - 3)\" using tpsL3_def tpsL_mvs' 4 by simp show ?thesis proof (cases "rs ! 2 < length M") case True then have "cmdL3 rs [!] j = (direction_to_symbol ((M ! (rs ! 2)) (drop (k + 3) rs) [~] (j - 3)), Stay)" using cmdL3_at_3a 4 drop_less_G by simp also have "... = (direction_to_symbol ((M ! (fst (exec t))) (config_read (exec t)) [~] (j - 3)), Stay)" using drop True rs2 by simp also have "... = (sim_move t ! (j - 3), Stay)" using sim_move_nth True 1 rs2 by simp finally have "cmdL3 rs [!] j = (sim_move t ! (j - 3), Stay)" . then show ?thesis using 2 3 act_onesie by (simp add: rs_def) next case False then have "cmdL3 rs [!] j = (1, Stay)" using cmdL3_at_3b 4 by simp then have "cmdL3 rs [!] j = (sim_move t ! (j - 3), Stay)" using sim_move_nth_else False 1 rs2 by simp then show ?thesis using 2 3 act_onesie by (simp add: rs_def) qed next case 5 then have 1: "j - k - 3 < k" by auto have 2: "tpsL2 t ! j = \exec t <.> (j - k - 3)\" using tpsL2_def tpsL_symbs' 5 by simp have 3: "tpsL3 t ! j = \sim_write t ! (j - k - 3)\" using tpsL3_def tpsL_symbs' 5 by simp show ?thesis proof (cases "rs ! 2 < length M") case True then have "cmdL3 rs [!] j = ((M ! (rs ! 2)) (drop (k + 3) rs) [.] (j - k - 3), Stay)" using 5 cmdL3_at_4a drop_less_G by simp also have "... = ((M ! (fst (exec t))) (config_read (exec t)) [.] (j - k - 3), Stay)" using drop True rs2 by simp also have "... = (sim_write t ! (j - k - 3), Stay)" using sim_write_nth 5 rs2 True by auto finally have "cmdL3 rs [!] j = (sim_write t ! (j - k - 3), Stay)" . then show ?thesis using 2 3 act_onesie rs_def by simp next case False then have "cmdL3 rs [!] j = (rs ! j, Stay)" using 5 cmdL3_at_4b by simp also have "... = (exec t <.> (j - k - 3), Stay)" using tpsL2_def read_tpsL_4 5 rs_def by simp also have "... = (config_read (exec t) ! (j - k - 3), Stay)" proof - have "length [k..drop (k + 3) rs = map (\j. exec t <.> j) [0.. drop by auto qed also have "... = (sim_write t ! (j - k - 3), Stay)" using sim_write_nth_else False rs2 by simp finally have "cmdL3 rs [!] j = (sim_write t ! (j - k - 3), Stay)" . then show ?thesis using 2 3 act_onesie rs_def by simp qed qed qed qed lemma execute_tmL2_3: "execute tmL2_3 (0, tpsL2 t) 1 = (1, tpsL3 t)" proof - have "execute tmL2_3 (0, tpsL2 t) 1 = exe tmL2_3 (execute tmL2_3 (0, tpsL2 t) 0)" by simp also have "... = exe tmL2_3 (0, tpsL2 t)" by simp also have "... = sem cmdL3 (0, tpsL2 t)" using tmL2_3_def by (simp add: exe_lt_length) finally show ?thesis using sem_cmdL3 by simp qed definition "esL3 t \ esL2 t @ [(n + 1, TT)]" lemma tmL3: assumes "t < TT" shows "traces tmL3 (tpsL0 t) (esL3 t) (tpsL3 t)" unfolding tmL3_def esL3_def proof (rule traces_sequential[OF tmL2[OF assms]]) show "traces tmL2_3 (tpsL2 t) [(n + 1, Suc (fmt n))] (tpsL3 t)" proof let ?es = "[(n + 1, TT)]" show "execute tmL2_3 (0, tpsL2 t) (length ?es) = (length tmL2_3, tpsL3 t)" using tmL2_3_def execute_tmL2_3 by simp show "\i. i < length ?es \ fst (execute tmL2_3 (0, tpsL2 t) i) < length tmL2_3" using tmL2_3_def execute_tmL2_3 by simp show "(execute tmL2_3 (0, tpsL2 t) (Suc i)) <#> 0 = fst (?es ! i) \ (execute tmL2_3 (0, tpsL2 t) (Suc i)) <#> 1 = snd (?es ! i)" if "i < length ?es" for i using execute_tmL2_3 that tpsL3_def tpsL_pos_0 tpsL_pos_1 by (metis One_nat_def fst_conv length_Cons less_one list.size(3) nth_Cons_0 snd_conv) qed qed definition "esL4 t \ esL3 t @ map (\i. (n + 1, i)) (rev [0.. tpsL t (replicate k (0, Some 0)) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" lemma tmL4: assumes "t < TT" shows "traces tmL4 (tpsL0 t) (esL4 t) (tpsL4 t)" unfolding tmL4_def esL4_def proof (rule traces_sequential[where ?tps2.0="tpsL3 t"]) show "traces tmL3 (tpsL0 t) (esL3 t) (tpsL3 t)" using tmL3 assms . show "traces tm_left_until1 (tpsL3 t) (map (Pair (n + 1)) (rev [0.. 1 < y \ dec y ! (2 * k + 1) = 1} (tpsL3 t ! 1)" using begin_tape_zip_cont tpsL3_def tpsL_def by simp show "map (Pair (n + 1)) (rev [0.. jj)" shows "enc_upd (zip_cont t xs i) (k + jj) 1 = zip_cont t (xs[jj:=(1, Some 1)]) i" proof - have "i < TT" using assms(1,4) exec_pos_less_TT by metis let ?n = "zip_cont t xs i" let ?xs = "xs[jj:=(1, Some 1)]" have "zip_cont t ?xs i = enc (map (\j. (exec (t + fst (?xs ! j)) <:> j) i) [0..j. case snd (?xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..i < TT\ dec_zip_cont) have lenys: "length ?ys = 2 * k + 2" by simp show "?ys ! j = (dec ?n) [k+jj:=1] ! j" if "j < length ?ys" for j proof - consider "j < k" | "k \ j \ j < 2 * k" | "j = 2 * k" | "j = 2 * k + 1" using lenys `j < length ?ys` by linarith then show ?thesis proof (cases) case 1 then have "?ys ! j = (exec (t + fst (?xs ! j)) <:> j) i" using assms(2) by (simp add: less_k_nth) have "(dec ?n) [k+jj:=1] ! j = dec ?n ! j" using 1 by simp also have "... = (exec (t + fst (xs ! j)) <:> j) i" by (simp add: "1" \i < TT\ dec_zip_cont_less_k) also have "... = (exec (t + fst (?xs ! j)) <:> j) i" using assms(1-3) by (metis fst_eqD nth_list_update) also have "... = ?ys ! j" using assms(2) 1 by (simp add: less_k_nth) finally show ?thesis by simp next case 2 show ?thesis proof (cases "j = k + jj") case True then have "?ys ! j = (case snd (?xs ! jj) of None \ 0 | Some d \ if i = exec (t + d) <#> jj then 1 else 0)" using assms(2) 2 by (simp add: less_2k_nth) also have "... = (if i = exec (t + 1) <#> jj then 1 else 0)" by (simp add: assms(1) assms(2)) also have "... = 1" using assms(1,4) by simp finally have "?ys ! j = 1" . moreover have "(dec ?n) [k+jj:=1] ! j = 1" using True len that by simp ultimately show ?thesis by simp next case False then have *: "?xs ! (j - k) = xs ! (j - k)" by (metis "2" add_diff_inverse_nat le_imp_less_Suc not_less_eq nth_list_update_neq) have "?ys ! j = (case snd (?xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using assms(2) 2 by (simp add: less_2k_nth) have "(dec ?n) [k+jj:=1] ! j = (dec ?n) ! j" using 2 False by simp also have "... = (case snd (xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using 2 `i < TT` dec_zip_cont_less_2k by simp also have "... = (case snd (?xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using * by simp also have "... = ?ys ! j" using assms(2) 2 by (simp add: less_2k_nth) finally show ?thesis by simp qed next case 3 then have "?ys ! j = (if i < t then 1 else 0)" by (simp add: twok_nth) moreover have "(dec ?n) [k+jj:=1] ! j = (if i < t then 1 else 0)" using 3 assms(1) dec_zip_cont_2k `i < TT` by simp ultimately show ?thesis by simp next case 4 then have "?ys ! j = (if i = 0 then 1 else 0)" using twok1_nth by fast moreover have "(dec ?n) [k+jj:=1] ! j = (if i = 0 then 1 else 0)" using 4 assms(1) dec_zip_cont_2k1 `i < TT` by simp ultimately show ?thesis by simp qed qed qed ultimately show ?thesis using enc_upd_def by simp qed lemma enc_upd_zip_cont_None_Some_Right: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, None)" and "i = Suc (exec t <#> jj)" and "sim_move t ! jj = 2" shows "enc_upd (zip_cont t xs i) (k + jj) 1 = zip_cont t (xs[jj:=(1, Some 1)]) i" proof - have "i = (exec (Suc t) <#> jj)" using assms sim_move by simp then show ?thesis using enc_upd_zip_cont_None_Some[OF assms(1-3)] by simp qed lemma enc_upd_zip_cont_None_Some_Left: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, None)" and "Suc i = exec t <#> jj" and "sim_move t ! jj = 0" shows "enc_upd (zip_cont t xs i) (k + jj) 1 = zip_cont t (xs[jj:=(1, Some 1)]) i" proof - have "i = (exec (Suc t) <#> jj)" using assms sim_move by simp then show ?thesis using enc_upd_zip_cont_None_Some[OF assms(1-3)] by simp qed lemma enc_upd_zip_cont_Some_None: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i = exec t <#> jj" shows "enc_upd (zip_cont t xs i) (k + jj) 0 = zip_cont t (xs[jj:=(1, None)]) i" proof - have "i < TT" using assms(1,4) by (simp add: exec_pos_less_TT) let ?n = "zip_cont t xs i" let ?xs = "xs[jj:=(1, None)]" have "zip_cont t ?xs i = enc (map (\j. (exec (t + fst (?xs ! j)) <:> j) i) [0..j. case snd (?xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..i < TT\ dec_zip_cont) have lenys: "length ?ys = 2 * k + 2" by simp show "?ys ! j = (dec ?n) [k+jj:=0] ! j" if "j < length ?ys" for j proof - consider "j < k" | "k \ j \ j < 2 * k" | "j = 2 * k" | "j = 2 * k + 1" using lenys `j < length ?ys` by linarith then show ?thesis proof (cases) case 1 then have "?ys ! j = (exec (t + fst (?xs ! j)) <:> j) i" using assms(2) by (simp add: less_k_nth) have "(dec ?n) [k+jj:=0] ! j = dec ?n ! j" using 1 by simp also have "... = (exec (t + fst (xs ! j)) <:> j) i" by (simp add: "1" \i < TT\ dec_zip_cont_less_k) also have "... = (exec (t + fst (?xs ! j)) <:> j) i" using assms(1-3) by (metis fst_eqD nth_list_update) also have "... = ?ys ! j" using assms(2) 1 by (simp add: less_k_nth) finally show ?thesis by simp next case 2 show ?thesis proof (cases "j = k + jj") case True then have "?ys ! j = (case snd (?xs ! jj) of None \ 0 | Some d \ if i = exec (t + d) <#> jj then 1 else 0)" using assms(2) 2 by (simp add: less_2k_nth) then have "?ys ! j = 0" using assms(1,2) by simp moreover have "(dec ?n) [k+jj:=0] ! j = 0" using True len that by simp ultimately show ?thesis by simp next case False then have *: "?xs ! (j - k) = xs ! (j - k)" by (metis "2" add_diff_inverse_nat le_imp_less_Suc not_less_eq nth_list_update_neq) have "?ys ! j = (case snd (?xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using assms(2) 2 by (simp add: less_2k_nth) have "(dec ?n) [k+jj:=0] ! j = (dec ?n) ! j" using 2 False by simp also have "... = (case snd (xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using 2 `i < TT` dec_zip_cont_less_2k by simp also have "... = (case snd (?xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using * by simp also have "... = ?ys ! j" using assms(2) 2 by (simp add: less_2k_nth) finally show ?thesis by simp qed next case 3 then have "?ys ! j = (if i < t then 1 else 0)" by (simp add: twok_nth) moreover have "(dec ?n) [k+jj:=0] ! j = (if i < t then 1 else 0)" using 3 assms(1) dec_zip_cont_2k `i < TT` by simp ultimately show ?thesis by simp next case 4 then have "?ys ! j = (if i = 0 then 1 else 0)" using twok1_nth by fast moreover have "(dec ?n) [k+jj:=0] ! j = (if i = 0 then 1 else 0)" using 4 assms(1) dec_zip_cont_2k1 `i < TT` by simp ultimately show ?thesis by simp qed qed qed ultimately show ?thesis using enc_upd_def by simp qed lemma zip_cont_nth_eq_updI1: assumes "i < TT" and "jj < k" and "length xs = k" and "xs ! jj = (0, Some 0)" and "(exec (Suc t) <:> jj) i = u" shows "enc_upd (zip_cont t xs i) jj u = zip_cont t (xs[jj:=(1, Some 0)]) i" proof - let ?n = "zip_cont t xs i" let ?xs = "xs[jj:=(1, Some 0)]" have "zip_cont t ?xs i = enc (map (\j. (exec (t + fst (?xs ! j)) <:> j) i) [0..j. case snd (?xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0.. j \ j < 2 * k" | "j = 2 * k" | "j = 2 * k + 1" using lenys `j < length ?ys` by linarith then show ?thesis proof (cases) case 1 then have *: "?ys ! j = (exec (t + fst (?xs ! j)) <:> j) i" by (simp add: less_k_nth) show ?thesis proof (cases "j = jj") case True then have "?ys ! j = (exec (Suc t) <:> j) i" using 1 assms(3) * by simp moreover have "(dec ?n) [jj:=u] ! j = u" using True len_eq that by simp ultimately show ?thesis using assms(5) True by simp next case False then have "(dec ?n) [jj:=u] ! j = (exec (t + fst (xs ! j)) <:> j) i" using dec_zip_cont_less_k 1 assms(1) by simp moreover have "?ys ! j = (exec (t + fst (xs ! j)) <:> j) i" using False * by simp ultimately show ?thesis by simp qed next case 2 then have *: "?ys ! j = (case snd (?xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" by (simp add: less_2k_nth) show ?thesis proof (cases "j = k + jj") case True then have "j - k = jj" by simp then have "snd (?xs ! (j - k)) = Some 0" using assms(2,3) by simp then have lhs: "?ys ! j = (if i = exec t <#> jj then 1 else 0)" using * True by simp have "(dec ?n) [jj:=u] ! j = (dec ?n) ! j" using True assms(2) by simp also have "... = (case snd (xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using True 2 assms(1,4) dec_zip_cont_less_2k by simp also have "... = (if i = exec t <#> jj then 1 else 0)" using True assms(4) by simp finally have "(dec ?n) [jj:=u] ! j = (if i = exec t <#> jj then 1 else 0)" . then show ?thesis using lhs True by simp next case False then have "j - k \ jj" using 2 by auto then have "snd (?xs ! (j - k)) = snd (xs ! (j - k))" by simp then have lhs: "?ys ! j = (case snd (xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using * by simp have "(dec ?n) [jj:=u] ! j = (dec ?n) ! j" using 2 assms(2) by simp then have "(dec ?n) [jj:=u] ! j = (case snd (xs ! (j - k)) of None \ 0 | Some d \ if i = exec (t + d) <#> (j - k) then 1 else 0)" using False 2 assms(1) dec_zip_cont_less_2k by simp then show ?thesis using lhs by simp qed next case 3 then have "?ys ! j = (if i < t then 1 else 0)" by (simp add: twok_nth) moreover have "(dec ?n) [jj:=u] ! j = (if i < t then 1 else 0)" using 3 assms(1,2) dec_zip_cont_2k by simp ultimately show ?thesis by simp next case 4 then have "?ys ! j = (if i = 0 then 1 else 0)" using twok1_nth by fast moreover have "(dec ?n) [jj:=u] ! j = (if i = 0 then 1 else 0)" using 4 assms(1,2) dec_zip_cont_2k1 by simp ultimately show ?thesis by simp qed qed qed ultimately show ?thesis using enc_upd_def by simp qed lemma zip_cont_upd_eq: assumes "jj < k" and "i = exec t <#> jj" and "i < TT" and "xs ! jj = (0, Some 0)" and "length xs = k" shows "(zip_cont t xs)(i:=enc_upd (zip_cont t xs i) jj (sim_write t ! jj)) = zip_cont t (xs[jj:=(1, Some 0)])" (is "?lhs = ?rhs") proof fix p consider "p < TT \ p \ i" | "p < TT \ p = i" | "p \ TT" by linarith then show "?lhs p = ?rhs p" proof (cases) case 1 then have "?lhs p = zip_cont t xs p" by simp moreover have "zip_cont t xs p = zip_cont t (xs[jj:=(1, Some 0)]) p" proof (rule zip_cont_nth_eqI) show "p < TT" using 1 by simp show "(exec (t + fst (xs ! j)) <:> j) p = (exec (t + fst (xs[jj := (1, Some 0)] ! j)) <:> j) p" if "j < k" for j proof (cases "j = jj") case True then have "fst (xs[jj := (1, Some 0)] ! j) = 1" using assms(1,5) by simp then have "(exec (t + fst (xs[jj := (1, Some 0)] ! j)) <:> j) p = (exec (Suc t) <:> j) p" by simp also have "... = (exec t <:> j) p" using assms(2) 1 by (simp add: True assms(1) sim_write) finally show ?thesis using assms(4) True by simp next case False then show ?thesis by simp qed show "snd (xs ! j) = snd (xs[jj := (1, Some 0)] ! j)" if "j < k" for j using assms(4,5) that by (cases "j = jj") simp_all qed ultimately show ?thesis by simp next case 2 then have "?lhs p = enc_upd (zip_cont t xs i) jj (sim_write t ! jj)" by simp then show ?thesis using 2 assms(1,2,4,5) sim_write' zip_cont_nth_eq_updI1 by auto next case 3 then show ?thesis using zip_cont_def assms(3) by auto qed qed lemma sem_cmdL5_neq_pos: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "snd (xs ! jj) = Some 0" and "i \ exec t <#> jj" and "i < TT" and "tps' = tpsL t xs (Suc i) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL5 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL5 jj)" using turing_command_cmdL5[OF assms(1)] turing_commandD(1) by simp show "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(6) by simp let ?rs = "read tps" have rs1: "?rs ! 1 \ \" using read_tpsL_1_bounds assms(2,5) by (metis not_one_less_zero) then show "fst (cmdL5 jj ?rs) = 0" by (simp add: cmdL5_def) show "act (cmdL5 jj ?rs [!] j) (tps ! j) = tps' ! j" if "j < 2 * k + 3" for j proof - consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_0 by simp then show ?thesis using assms tpsL_0 1 by (metis act_Stay that length_tpsL) next case 2 then have "enc_nth (?rs ! 1) (k + jj) = (if i = exec t <#> jj then 1 else 0)" using assms read_tpsL_1_nth_less_2k by simp then have "enc_nth (?rs ! 1) (k + jj) = 0" using assms(4) by simp then have "\ (1 < ?rs ! 1 \ ?rs ! 1 < G^(2*k+2)+2 \ ?rs ! (3+k+jj) < G \ enc_nth (?rs!1) (k+jj) = 1)" by simp then have "cmdL5 jj ?rs [!] 1 = (?rs ! 1, Right)" using cmdL5_at_1_else rs1 by simp then show ?thesis using assms tpsL_1 2 act_Right that length_tpsL by (metis Suc_eq_plus1 prod.sel(1) tpsL_pos_1) next case 3 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_2 by simp then show ?thesis using assms tpsL_2 3 by (metis act_Stay that length_tpsL) next case 4 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_3 by simp then have "act (cmdL5 jj ?rs [!] j) (tps ! j) = tps ! j" using act_Stay by (simp add: \length tps = 2 * k + 3\ that) then show ?thesis using assms tpsL_mvs' by (simp add: "4" add.commute) next case 5 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_4 by simp then have "act (cmdL5 jj ?rs [!] j) (tps ! j) = tps ! j" using act_Stay by (simp add: \length tps = 2 * k + 3\ that) then show ?thesis using assms tpsL_symbs' 5 by simp qed qed qed lemma sem_cmdL5_eq_pos: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "xs ! jj = (0, Some 0)" and "i = exec t <#> jj" and "tps' = tpsL t (xs[jj:=(1, Some 0)]) (Suc i) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" shows "sem (cmdL5 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) have "i < TT" using exec_pos_less_TT assms(1,4) by simp show "proper_command (2 * k + 3) (cmdL5 jj)" using turing_command_cmdL5[OF assms(1)] turing_commandD(1) by simp show "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(5) by simp let ?rs = "read tps" have rs1: "?rs ! 1 \ \" using read_tpsL_1_bounds assms(2) `i < TT` by (metis not_one_less_zero) then show "fst (cmdL5 jj ?rs) = 0" by (simp add: cmdL5_def) show "act (cmdL5 jj (read tps) [!] j) (tps ! j) = tps' ! j" if "j < 2 * k + 3" for j proof - consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_0 by simp then show ?thesis using assms tpsL_0 1 by (metis act_Stay that length_tpsL) next case 2 then have "enc_nth (?rs ! 1) (k + jj) = (if i = exec t <#> jj then 1 else 0)" using assms `i < TT` read_tpsL_1_nth_less_2k by simp then have "enc_nth (?rs ! 1) (k + jj) = 1" using assms(4) by simp moreover have "1 < ?rs ! 1 \ ?rs ! 1 < G^(2*k+2)+2" using assms(2) `i < TT` read_tpsL_1_bounds by auto moreover have "?rs ! (3+k+jj) < G" using read_tpsL_4 assms sim_write_nth_less_G[OF assms(1)] by simp ultimately have "1 < ?rs ! 1 \ ?rs ! 1 < G^(2*k+2)+2 \ ?rs ! (3+k+jj) < G \ enc_nth (?rs!1) (k+jj) = 1" by simp then have "cmdL5 jj ?rs [!] 1 = (enc_upd (?rs!1) jj (?rs!(3+k+jj)), Right)" using cmdL5_at_1 rs1 by simp moreover have "?rs!(3+k+jj) = sim_write t ! jj" by (simp add: assms(1,2) read_tpsL_4) ultimately have *: "cmdL5 jj ?rs [!] 1 = (enc_upd (?rs ! 1) jj (sim_write t ! jj), Right)" by simp have "?rs ! 1 = zip_cont t xs i" using assms(2) read_tpsL_1 zip_cont_def by auto let ?tp = "act (enc_upd (?rs ! 1) jj (sim_write t ! jj), Right) (tps ! j)" have "?tp = tps' ! 1" proof - have "?tp = ((zip_cont t xs)(i:=enc_upd (?rs ! 1) jj (sim_write t ! jj)), Suc i)" using act_Right' assms tpsL_1 by (metis "2" add.commute fst_conv plus_1_eq_Suc snd_conv) moreover have "tps' ! 1 = (zip_cont t (xs[jj:=(1, Some 0)]), Suc i)" using assms(5) tpsL_1 by simp moreover have "(zip_cont t xs)(i:=enc_upd (?rs ! 1) jj (sim_write t ! jj)) = zip_cont t (xs[jj:=(1, Some 0)])" using zip_cont_upd_eq assms `i < TT` \read tps ! 1 = zip_cont t xs i\ by auto ultimately show ?thesis by auto qed then show ?thesis using 2 * by simp next case 3 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_2 by simp then show ?thesis using assms tpsL_2 3 by (metis act_Stay that length_tpsL) next case 4 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_3 by simp then have "act (cmdL5 jj ?rs [!] j) (tps ! j) = tps ! j" using act_Stay by (simp add: \length tps = 2 * k + 3\ that) then show ?thesis using assms tpsL_mvs' by (simp add: "4" add.commute) next case 5 then have "cmdL5 jj ?rs [!] j = (?rs ! j, Stay)" using rs1 cmdL5_at_4 by simp then have "act (cmdL5 jj ?rs [!] j) (tps ! j) = tps ! j" using act_Stay by (simp add: \length tps = 2 * k + 3\ that) then show ?thesis using assms tpsL_symbs' 5 by simp qed qed qed lemma sem_cmdL5_eq_TT: assumes "jj < k" and "tps = tpsL t xs TT q mvs symbs" shows "sem (cmdL5 jj) (0, tps) = (1, tps)" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL5 jj)" using turing_command_cmdL5[OF assms(1)] turing_commandD(1) by simp show "length tps = 2 * k + 3" using assms(2) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp let ?rs = "read tps" have rs1: "?rs ! 1 = \" using read_tpsL_1 assms(2) by simp then show "fst (cmdL5 jj ?rs) = 1" by (simp add: cmdL5_def) show "\i. i < 2 * k + 3 \ act (cmdL5 jj ?rs [!] i) (tps ! i) = tps ! i" using len rs1 act_Stay cmdL5_eq_0 read_length by auto qed lemma execute_tmL45_1: assumes "tt \ exec t <#> jj" and "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL45 jj) (0, tps) tt = (0, tps')" using assms(1,5) proof (induction tt arbitrary: tps') case 0 then show ?case using assms(2-4) by simp next case (Suc tt) then have tt_neq: "tt \ exec t <#> jj" by simp have tt_le: "tt \ exec t <#> jj" using Suc.prems by simp then have tt_less: "tt < TT" using exec_pos_less_TT assms(2) by (meson le_trans less_Suc_eq_le) define tps_tt where "tps_tt = tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" have "execute (tmL45 jj) (0, tps) (Suc tt) = exe (tmL45 jj) (execute (tmL45 jj) (0, tps) tt)" by simp also have "... = exe (tmL45 jj) (0, tps_tt)" using Suc.IH assms(2-4) tt_le tps_tt_def by simp also have "... = sem (cmdL5 jj) (0, tps_tt)" using tmL45_def exe_lt_length by simp also have "... = (0, tpsL t xs (Suc tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using sem_cmdL5_neq_pos assms(2-4) tt_neq tt_less by (simp add: tps_tt_def) finally show ?case by (simp add: Suc.prems(2)) qed lemma execute_tmL45_2: assumes "tt = exec t <#> jj" and "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t (xs[jj:=(1, Some 0)]) (Suc tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" shows "execute (tmL45 jj) (0, tps) (Suc tt) = (0, tps')" proof - have "execute (tmL45 jj) (0, tps) (Suc tt) = exe (tmL45 jj) (execute (tmL45 jj) (0, tps) tt)" by simp also have "... = exe (tmL45 jj) (0, tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using execute_tmL45_1 assms by simp also have "... = sem (cmdL5 jj) (0, tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using tmL45_def exe_lt_length by simp also have "... = (0, tpsL t (xs[jj:=(1, Some 0)]) (Suc tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using sem_cmdL5_eq_pos[OF assms(2)] assms by simp finally show ?thesis using assms(5) by simp qed lemma execute_tmL45_3: assumes "tt \ Suc (exec t <#> jj)" and "tt \ TT" and "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t (xs[jj:=(1, Some 0)]) tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" shows "execute (tmL45 jj) (0, tps) tt = (0, tps')" using assms(1,2,6) proof (induction tt arbitrary: tps' rule: nat_induct_at_least) case base then show ?case using assms(3-5,7) execute_tmL45_2 by simp next case (Suc tt) then have tt: "tt < TT" "tt \ exec t <#> jj" by simp_all have "execute (tmL45 jj) (0, tps) (Suc tt) = exe (tmL45 jj) (execute (tmL45 jj) (0, tps) tt)" by simp also have "... = exe (tmL45 jj) (0, tpsL t (xs[jj:=(1, Some 0)]) tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using Suc by simp also have "... = sem (cmdL5 jj) (0, tpsL t (xs[jj:=(1, Some 0)]) tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using tmL45_def exe_lt_length by simp also have "... = (0, tpsL t (xs[jj:=(1, Some 0)]) (Suc tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using sem_cmdL5_neq_pos tt by (simp add: assms(3) assms(7)) finally show ?case using Suc(4) by presburger qed lemma execute_tmL45_4: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t (xs[jj:=(1, Some 0)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" shows "execute (tmL45 jj) (0, tps) (Suc TT) = (1, tps')" proof - have "execute (tmL45 jj) (0, tps) (Suc TT) = exe (tmL45 jj) (execute (tmL45 jj) (0, tps) TT)" by simp also have "... = exe (tmL45 jj) (0, tps')" using assms execute_tmL45_3 by (metis Suc_leI exec_pos_less_TT order_refl) also have "... = sem (cmdL5 jj) (0, tps')" using tmL45_def exe_lt_length by simp also have "... = (1, tps')" using sem_cmdL5_eq_TT assms(1,4) by simp finally show ?thesis . qed definition "esL45 \ map (\i. (n + 1, Suc i)) [0..j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t (xs[jj:=(1, Some 0)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL45 jj) tps esL45 tps'" proof show "execute (tmL45 jj) (0, tps) (length esL45) = (length (tmL45 jj), tps')" using tmL45_def execute_tmL45_4 esL45_def assms by simp show "fst (execute (tmL45 jj) (0, tps) i) < length (tmL45 jj)" if "i < length esL45" for i proof - have "i \ TT" using esL45_def that by simp then consider "i \ exec t <#> jj" | "i = Suc (exec t <#> jj)" | "i > Suc (exec t <#> jj) \ i \ TT" by linarith then show ?thesis proof (cases) case 1 then show ?thesis using assms execute_tmL45_1 tmL45_def by simp next case 2 then show ?thesis using assms execute_tmL45_2 tmL45_def by simp next case 3 then show ?thesis using assms execute_tmL45_3 tmL45_def by simp qed qed show "execute (tmL45 jj) (0, tps) (Suc i) <#> 0 = fst (esL45 ! i) \ execute (tmL45 jj) (0, tps) (Suc i) <#> 1 = snd (esL45 ! i)" if "i < length esL45" for i proof - have "i \ TT" using esL45_def that by simp then consider "i < exec t <#> jj" | "i = exec t <#> jj" | "i \ Suc (exec t <#> jj) \ i < TT" | "i = TT" by linarith then show ?thesis proof (cases) case 1 then have "Suc i \ exec t <#> jj" by simp then have "i < TT" using exec_pos_less_TT by (metis \i \ TT\ assms(1) nat_less_le not_less_eq_eq) then have "esL45 ! i = (n + 1, Suc i)" using esL45_def nth_map_upt_TT by auto then show ?thesis using assms execute_tmL45_1 tpsL_pos_0 tpsL_pos_1 by (metis \Suc i \ exec t <#> jj\ fst_conv snd_conv) next case 2 then have "i < TT" using exec_pos_less_TT by (simp add: assms(1)) then have "esL45 ! i = (n + 1, Suc i)" using esL45_def nth_map_upt_TT by auto then show ?thesis using assms execute_tmL45_2 tpsL_pos_0 tpsL_pos_1 by (metis 2 fst_conv snd_conv) next case 3 then have "esL45 ! i = (n + 1, Suc i)" using esL45_def nth_map_upt_TT by auto then show ?thesis using assms execute_tmL45_3 tpsL_pos_0 tpsL_pos_1 3 by (metis Suc_leI fst_conv le_imp_less_Suc nat_less_le snd_conv) next case 4 then have "esL45 ! i = (n + 1, TT)" using esL45_def by (simp add: nth_append) then show ?thesis using assms execute_tmL45_4 tpsL_pos_0 tpsL_pos_1 4 by simp qed qed qed definition "esL46 \ esL45 @ [(n + 1, fmt n)]" lemma len_esL46: "length esL46 = TT + 2" using len_esL45 esL46_def by simp lemma tmL46: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t (xs[jj:=(1, Some 0)]) (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL46 jj) tps esL46 tps'" unfolding tmL46_def esL46_def proof (rule traces_sequential) let ?tps = "tpsL t (xs[jj:=(1, Some 0)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" show "traces (tmL45 jj) tps esL45 ?tps" using tmL45 assms by simp show "traces (tm_left 1) ?tps [(n + 1, fmt n)] tps'" using tpsL_pos_0 tpsL_pos_1 assms tpsL_def tpsL_1 by (intro traces_tm_left_1I) simp_all qed lemma sem_cmdL7_nonleft_gt_0: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "i < TT" and "i > 0" and "sim_move t ! jj \ 0" and "tps' = tpsL t xs (i - 1) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL7 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL7 jj)" using turing_command_cmdL7[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(7) by simp define rs where "rs = read tps" then have "\ is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms by (metis enc_nth_dec nat_neq_iff numerals(1) zero_neq_numeral) then show "fst (cmdL7 jj rs) = 0" unfolding cmdL7_def by simp have "rs ! (3 + jj) = sim_move t ! jj" using rs_def assms(1,2) read_tpsL_3 by simp moreover have "sim_move t ! jj < 3" using sim_move_def assms(1) direction_to_symbol_less sim_move_nth sim_move_nth_else by (metis One_nat_def not_add_less2 not_less_eq numeral_3_eq_3 plus_1_eq_Suc) ultimately have "condition7c rs jj" using assms(6) by simp then have *: "snd (cmdL7 jj rs) = [(rs ! 0, Stay), (rs ! 1, Left), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,7) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (rs ! 1, Left) (tps ! 1)" using * rs_def by simp also have "... = tps' ! 1" using act_Left len rs_def assms tpsL_1 by (metis 2 fst_conv that tpsL_pos_1) also have "... = tps' ! j" using 2 by simp finally show ?thesis . next case 3 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,7) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_mvs' 4 by simp finally show ?thesis . next case 5 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma sem_cmdL7_nonleft_eq_0: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "sim_move t ! jj \ 0" shows "sem (cmdL7 jj) (0, tps) = (1, tps)" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL7 jj)" using turing_command_cmdL7[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps = 2 * k + 3" using assms(2) by simp define rs where "rs = read tps" then have "is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms enc_nth_def read_tpsL_1_bounds zero_less_Suc by simp then show "fst (cmdL7 jj (read tps)) = 1" using cmdL7_def rs_def by simp have "rs ! (3 + jj) = sim_move t ! jj" using rs_def assms(1,2) read_tpsL_3 by simp then have "condition7c rs jj" using sim_move direction_to_symbol_less sim_move_nth sim_move_nth_else assms(1,4) by (metis less_Suc_eq not_add_less2 numeral_3_eq_3 numeral_eq_iff numerals(1) plus_1_eq_Suc semiring_norm(86)) then have *: "snd (cmdL7 jj rs) = [(rs ! 0, Stay), (rs ! 1, Left), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then show ?thesis using * act_Stay assms(2) rs_def by simp next case 2 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! 1, Left) (tps ! j)" using * rs_def by simp also have "... = tps ! j" using 2 assms(2) act_Left that length_tpsL tpsL_1 tpsL_pos_1 rs_def by (metis diff_is_0_eq' fst_conv less_numeral_extra(1) nat_less_le) finally show ?thesis by simp next case 3 then show ?thesis using * act_Stay assms(2) rs_def by simp next case 4 then show ?thesis using * act_Stay assms rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] len by simp next case 5 then show ?thesis using * act_Stay assms rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] len by simp qed qed qed lemma execute_tmL67_nonleft_less: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "sim_move t ! jj \ 0" and "tt < TT" and "tps' = tpsL t xs (fmt n - tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL67 jj) (0, tps) tt = (0, tps')" using assms(5,6) proof (induction tt arbitrary: tps') case 0 then show ?case using assms(1-4) tmL67_def by simp next case (Suc tt) have "execute (tmL67 jj) (0, tps) (Suc tt) = exe (tmL67 jj) (execute (tmL67 jj) (0, tps) tt)" by simp also have "... = exe (tmL67 jj) (0, tpsL t xs (fmt n - tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using Suc by simp finally show ?case using assms(1-4) sem_cmdL7_nonleft_gt_0 tmL67_def exe_lt_length Suc by simp qed lemma execute_tmL67_nonleft_finish: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "sim_move t ! jj \ 0" and "tps' = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL67 jj) (0, tps) TT = (1, tps')" using assms execute_tmL67_nonleft_less sem_cmdL7_nonleft_eq_0 tmL67_def exe_lt_length by simp definition "esL67 \ map (\i. (n + 1, i)) (rev [0.. esL67 ! i = (n + 1, fmt n - i - 1)" proof - assume "i < fmt n" then have "(rev [0..i < fmt n\ esL67_def nth_append) qed lemma tmL67_nonleft: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "sim_move t ! jj \ 0" and "tps' = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL67 jj) tps esL67 tps'" proof have len: "length esL67 = TT" using esL67_def by simp then show 1: "execute (tmL67 jj) (0, tps) (length esL67) = (length (tmL67 jj), tps')" using assms tmL67_def execute_tmL67_nonleft_finish by simp show "\i. i < length esL67 \ fst (execute (tmL67 jj) (0, tps) i) < length (tmL67 jj)" using len assms execute_tmL67_nonleft_less tmL67_def by simp show "(execute (tmL67 jj) (0, tps) (Suc i)) <#> 0 = fst (esL67 ! i) \ (execute (tmL67 jj) (0, tps) (Suc i)) <#> 1 = snd (esL67 ! i)" if "i < length esL67" for i proof (cases "i = fmt n") case True then show ?thesis using assms that 1 tpsL_pos_0 tpsL_pos_1 len by simp next case False then have "Suc i < TT" using that len by simp moreover from this have "esL67 ! i = (n + 1, fmt n - 1 - i)" by simp ultimately show ?thesis using tpsL_pos_0 tpsL_pos_1 assms(1-5) execute_tmL67_nonleft_less by (metis (no_types, lifting) diff_diff_left fst_conv plus_1_eq_Suc snd_conv) qed qed lemma sem_cmdL7_1: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i < TT" and "i > exec t <#> jj" and "sim_move t ! jj = 0" and "tps' = tpsL t xs (i - 1) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL7 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL7 jj)" using turing_command_cmdL7[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(8) by simp define rs where "rs = read tps" then have not_beginning: "\ is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms enc_nth_def read_tpsL_1_bounds zero_less_Suc by simp then show "fst (cmdL7 jj (read tps)) = 0" using cmdL7_def rs_def by simp have "rs ! (3 + jj) = \" using rs_def read_tpsL_3 assms by simp moreover have "enc_nth (rs ! 1) (k + jj) = 0" using assms rs_def read_tpsL_1_nth_less_2k by simp ultimately have "condition7c rs jj" using not_beginning by simp then have *: "snd (cmdL7 jj rs) = [(rs ! 0, Stay), (rs ! 1, Left), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,8) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (rs ! 1, Left) (tps ! 1)" using * rs_def by simp also have "... = tps' ! 1" using act_Left len rs_def assms tpsL_1 by (metis 2 fst_conv that tpsL_pos_1) also have "... = tps' ! j" using 2 by simp finally show ?thesis . next case 3 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,8) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,8) tpsL_mvs' 4 by simp finally show ?thesis . next case 5 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,8) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL67_1: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "tt < TT - exec t <#> jj" and "tps' = tpsL t xs (fmt n - tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL67 jj) (0, tps) tt = (0, tps')" using assms(6,7) proof (induction tt arbitrary: tps') case 0 then show ?case by (simp add: assms(2)) next case (Suc tt) then have "execute (tmL67 jj) (0, tps) (Suc tt) = sem (cmdL7 jj) (execute (tmL67 jj) (0, tps) tt)" using exe_lt_length tmL67_def by simp then show ?case using assms(1-5) sem_cmdL7_1 Suc by simp qed lemma zip_cont_enc_upd_Some: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, None)" and "i = exec (Suc t) <#> jj" shows "(zip_cont t xs)(i:=(enc_upd (zip_cont t xs i) (k + jj) 1)) = zip_cont t (xs[jj:=(1, Some 1)])" (is "?lhs = ?rhs") proof fix p have "i < TT" using assms(1,4) exec_pos_less_TT by simp consider "p < TT \ p \ i" | "p < TT \ p = i" | "p \ TT" by linarith then show "?lhs p = ?rhs p" proof (cases) case 1 then have "?lhs p = zip_cont t xs p" by simp moreover have "zip_cont t xs p = zip_cont t (xs[jj:=(1, Some 1)]) p" proof (rule zip_cont_eqI) show "p < TT" using 1 by simp show "(exec (t + fst (xs ! j)) <:> j) p = (exec (t + fst (xs[jj := (1, Some 1)] ! j)) <:> j) p" if "j < k" for j using assms(1-3) by (cases "j = jj") simp_all show "(case snd (xs ! j) of None \ 0 | Some d \ if p = (exec (t + d)) <#> j then 1 else 0) = (case snd (xs[jj := (1, Some 1)] ! j) of None \ 0 | Some d \ if p = (exec (t + d)) <#> j then 1 else 0)" (is "?lhs = ?rhs") if "j < k" for j proof (cases "j = jj") case True then show ?thesis using 1 assms by simp next case False then show ?thesis by simp qed qed ultimately show ?thesis by simp next case 2 then show ?thesis using assms enc_upd_zip_cont_None_Some by simp next case 3 then show ?thesis using `i < TT` zip_cont_eq_0 by simp qed qed lemma zip_cont_enc_upd_Some_Right: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, None)" and "i = Suc (exec t <#> jj)" and "sim_move t ! jj = 2" shows "(zip_cont t xs)(i:=(enc_upd (zip_cont t xs i) (k + jj) 1)) = zip_cont t (xs[jj:=(1, Some 1)])" proof - have "i = exec (Suc t) <#> jj" using assms sim_move by simp then show ?thesis using zip_cont_enc_upd_Some[OF assms(1-3)] by simp qed lemma zip_cont_enc_upd_Some_Left: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, None)" and "Suc i = exec t <#> jj" and "sim_move t ! jj = 0" shows "(zip_cont t xs)(i:=(enc_upd (zip_cont t xs i) (k + jj) 1)) = zip_cont t (xs[jj:=(1, Some 1)])" (is "?lhs = ?rhs") proof - have "i = exec (Suc t) <#> jj" using assms sim_move by simp then show ?thesis using zip_cont_enc_upd_Some[OF assms(1-3)] by simp qed lemma zip_cont_enc_upd_None: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i = exec t <#> jj" shows "(zip_cont t xs)(i:=(enc_upd (zip_cont t xs i) (k + jj) 0)) = zip_cont t (xs[jj:=(1, None)])" (is "?lhs = ?rhs") proof fix p consider "p < TT \ p \ i" | "p < TT \ p = i" | "p \ TT" by linarith then show "?lhs p = ?rhs p" proof (cases) case 1 then have "?lhs p = zip_cont t xs p" by simp moreover have "zip_cont t xs p = zip_cont t (xs[jj:=(1, None)]) p" proof (rule zip_cont_eqI) show "p < TT" using 1 by simp show "(exec (t + fst (xs ! j)) <:> j) p = (exec (t + fst (xs[jj := (1, None)] ! j)) <:> j) p" if "j < k" for j using assms(1-3) by (cases "j = jj") simp_all show "(case snd (xs ! j) of None \ 0 | Some d \ if p = (exec (t + d)) <#> j then 1 else 0) = (case snd (xs[jj := (1, None)] ! j) of None \ 0 | Some d \ if p = (exec (t + d)) <#> j then 1 else 0)" if "j < k" for j using assms 1 by (cases "j = jj") simp_all qed ultimately show ?thesis by simp next case 2 then have "?lhs p = enc_upd (zip_cont t xs i) (k + jj) 0" by simp moreover have "enc_upd (zip_cont t xs i) (k + jj) 0 = zip_cont t (xs[jj:=(1, None)]) i" using assms(1-4) enc_upd_zip_cont_Some_None by simp ultimately show ?thesis using 2 by simp next case 3 moreover have "i < TT" using assms(4) by (simp add: assms(1) exec_pos_less_TT) ultimately show ?thesis using zip_cont_eq_0 by simp qed qed lemma sem_cmdL7_2a: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i = exec t <#> jj" and "i > 0" and "sim_move t ! jj = 0" and "tps' = tpsL t (xs[jj:=(1, None)]) (i - 1) 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL7 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL7 jj)" using turing_command_cmdL7[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(8) by simp define rs where "rs = read tps" then have not_beginning: "\ is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms enc_nth_def read_tpsL_1_bounds zero_less_Suc exec_pos_less_TT by simp then show "fst (cmdL7 jj (read tps)) = 0" using cmdL7_def rs_def by simp have "i < TT" using assms(5) by (simp add: assms(1) exec_pos_less_TT) have "rs ! (3 + jj) = \" using rs_def read_tpsL_3 assms by simp moreover have "enc_nth (rs ! 1) (k + jj) = 1" using assms rs_def read_tpsL_1_nth_less_2k[OF `i < TT`] by simp ultimately have "condition7a rs jj" using not_beginning `i < TT` assms(2) read_tpsL_1_bounds rs_def by simp then have *: "snd (cmdL7 jj rs) = [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 0, Left), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 3 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,8) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (enc_upd (rs ! 1) (k + jj) 0, Left) (tps ! 1)" using * rs_def by simp also have "... = tps ! 1 |:=| (enc_upd (rs ! 1) (k + jj) 0) |-| 1" using act_Left' 2 len by simp also have "... = tps' ! 1" using assms zip_cont_enc_upd_None rs_def read_tpsL_1 tpsL_1 zip_cont_def by simp finally show ?thesis using 2 by simp next case 3 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,8) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 show ?thesis proof (cases "j = 3 + jj") case True then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (3, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by (smt (verit)) also have "... = tps' ! j" using 4 assms(2,8) True act_onesie tpsL_mvs by simp finally show ?thesis . next case False then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by auto also have "... = tps' ! j" using 4 assms(2,8) False act_Stay len rs_def that tpsL_mvs' - by (smt (z3) add.commute le_add_diff_inverse2) + by (smt (verit) add.commute le_add_diff_inverse2) finally show ?thesis . qed next case 5 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,8) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL67_2a: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "exec t <#> jj > 0" and "tt = TT - exec t <#> jj" and "tps' = tpsL t (xs[jj:=(1, None)]) (fmt n - tt) 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL67 jj) (0, tps) tt = (0, tps')" proof - have "tt > 0" using assms(1,7) exec_pos_less_TT by simp then have "tt - 1 < TT - exec t <#> jj" using assms(6,7) by simp then have *: "execute (tmL67 jj) (0, tps) (tt - 1) = (0, tpsL t xs (fmt n - tt + 1) 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using assms execute_tmL67_1[where ?tt="tt - 1"] by simp have **: "fmt n - tt + 1 = exec t <#> jj" - using assms(1,6,7) exec_pos_less_TT - by (smt (z3) Nat.add_diff_assoc2 Suc_diff_Suc add.right_neutral add_Suc_right add_diff_cancel_right' - diff_Suc_Suc diff_less le_add_diff_inverse2 nat_less_le plus_1_eq_Suc zero_less_Suc) + using assms(1,6,7) exec_pos_less_TT Suc_diff_le less_eq_Suc_le by auto have "execute (tmL67 jj) (0, tps) tt = exe (tmL67 jj) (execute (tmL67 jj) (0, tps) (tt - 1))" using `tt > 0` exe_lt_length by (metis One_nat_def Suc_diff_Suc diff_zero execute.simps(2)) also have "... = sem (cmdL7 jj) (execute (tmL67 jj) (0, tps) (tt - 1))" using tmL67_def exe_lt_length * by simp also have "... = sem (cmdL7 jj) (0, tpsL t xs (fmt n - tt + 1) 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using * by simp also have "... = (0, tpsL t (xs[jj:=(1, None)]) (fmt n - tt) 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j))" using ** assms sem_cmdL7_2a[where ?i="fmt n - tt + 1"] by simp finally show ?thesis using assms(8) by simp qed lemma zip_cont_Some_Some: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i = exec t <#> jj" and "i = 0" and "sim_move t ! jj = 0" shows "zip_cont t xs = zip_cont t (xs[jj:=(1, Some 1)])" (is "?lhs = ?rhs") proof fix p consider "p < TT" | "p \ TT" by linarith then show "?lhs p = ?rhs p" proof (cases) case 1 then have "?lhs p = zip_cont t xs p" by simp moreover have "zip_cont t xs p = zip_cont t (xs[jj:=(1, Some 1)]) p" proof (rule zip_cont_eqI) show "p < TT" using 1 by simp show "\j. j < k \ ((exec (t + fst (xs ! j))) <:> j) p = ((exec (t + fst (xs[jj := (1, Some 1)] ! j))) <:> j) p" by (metis assms(2,3) fst_conv nth_list_update_eq nth_list_update_neq) show "\j. j < k \ (case snd (xs ! j) of None \ 0 | Some d \ if p = exec (t + d) <#> j then 1 else 0) = (case snd (xs[jj := (1, Some 1)] ! j) of None \ 0 | Some d \ if p = exec (t + d) <#> j then 1 else 0)" using assms 1 sim_move by (metis (no_types, lifting) add.commute add.right_neutral diff_add_zero nth_list_update_eq nth_list_update_neq option.simps(5) plus_1_eq_Suc prod.sel(2)) qed ultimately show ?thesis by simp next case 2 then show ?thesis using zip_cont_eq_0 by simp qed qed lemma sem_cmdL7_2b: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i = exec t <#> jj" and "i = 0" and "sim_move t ! jj = 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL7 jj) (0, tps) = (1, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL7 jj)" using turing_command_cmdL7[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(8) by simp define rs where "rs = read tps" then have is_beginning: "is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms(2,6) enc_nth_def read_tpsL_1_bounds rs_def by simp then show "fst (cmdL7 jj (read tps)) = 1" using assms(6) cmdL7_def rs_def by simp have "rs ! (3 + jj) = \" by (simp add: rs_def assms add.commute read_tpsL_3') then have "condition7c rs jj" using is_beginning by simp then have *: "snd (cmdL7 jj rs) = [(rs ! 0, Stay), (rs ! 1, Left), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,8) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (rs ! 1, Left) (tps ! 1)" using * rs_def by simp also have "... = tps' ! 1" using zip_cont_Some_Some assms rs_def tpsL_1 "2" act_Left fst_conv len that tpsL_pos_1 by (metis zero_diff) finally show ?thesis using 2 by simp next case 3 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,8) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by auto also have "... = tps' ! j" using 4 assms(2,8) act_Stay len rs_def that tpsL_mvs' by (metis add.commute) finally show ?thesis . next case 5 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,8) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL67_2b: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "exec t <#> jj = 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL67 jj) (0, tps) TT = (1, tps')" using execute_tmL67_1 assms exe_lt_length tmL67_def sem_cmdL7_2b by simp lemma tmL67_left_0: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "exec t <#> jj = 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL67 jj) tps esL67 tps'" proof show "execute (tmL67 jj) (0, tps) (length esL67) = (length (tmL67 jj), tps')" using esL67_def tmL67_def execute_tmL67_2b assms by simp show "\i. i < length esL67 \ fst (execute (tmL67 jj) (0, tps) i) < length (tmL67 jj)" using esL67_def tmL67_def execute_tmL67_1 assms by simp show "execute (tmL67 jj) (0, tps) (Suc i) <#> 0 = fst (esL67 ! i) \ execute (tmL67 jj) (0, tps) (Suc i) <#> 1 = snd (esL67 ! i)" if "i < length esL67" for i proof (cases "i = fmt n") case True then have "Suc i = TT" by simp moreover have "esL67 ! i = (n + 1, 0)" using True esL67_def by (simp add: nth_append) ultimately show ?thesis using assms that tpsL_pos_0 tpsL_pos_1 by (metis execute_tmL67_2b fst_conv snd_conv) next case False then have "Suc i < TT" using that esL67_def by simp moreover from this have "esL67 ! i = (n + 1, fmt n - 1 - i)" by simp ultimately show ?thesis using tpsL_pos_0 tpsL_pos_1 assms(1-6) execute_tmL67_1 by (metis (no_types, lifting) diff_diff_left fst_conv minus_nat.diff_0 plus_1_eq_Suc snd_conv) qed qed lemma sem_cmdL7_3: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, None)" and "Suc i = exec t <#> jj" and "sim_move t ! jj = 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) (i - 1) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL7 jj) (0, tps) = (if i = 0 then 1 else 0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL7 jj)" using turing_command_cmdL7[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(7) by simp define rs where "rs = read tps" show "fst (cmdL7 jj (read tps)) = (if i = 0 then 1 else 0)" proof (cases "i = 0") case True then have "is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms(2) enc_nth_def read_tpsL_1_bounds rs_def by simp then show ?thesis using True cmdL7_def rs_def by simp next case False then have "\ is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms enc_nth_def exec_pos_less_TT by (metis (no_types, lifting) Suc_le_lessD less_imp_le_nat less_numeral_extra(1) neq0_conv rs_def) then show ?thesis using False cmdL7_def rs_def by simp qed have "i < TT" using assms exec_pos_less_TT by (metis Suc_less_eq less_SucI) have "rs ! (3 + jj) = 3" by (simp add: rs_def assms(1,2) add.commute read_tpsL_3') moreover have "is_code (rs ! 1)" using assms rs_def read_tpsL_1_nth_less_2k `i < TT` read_tpsL_1_bounds by simp ultimately have "condition7b rs jj" using `i < TT` assms(2) read_tpsL_1_bounds rs_def by simp then have *: "snd (cmdL7 jj rs) = [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 1, Left), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 0 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,7) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (enc_upd (rs ! 1) (k + jj) 1, Left) (tps ! 1)" using * rs_def by simp also have "... = tps ! 1 |:=| (enc_upd (rs ! 1) (k + jj) 1) |-| 1" using act_Left' 2 len by simp also have "... = tps' ! 1" using assms zip_cont_enc_upd_Some_Left rs_def read_tpsL_1 tpsL_1 zip_cont_def by simp finally show ?thesis using 2 by simp next case 3 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,7) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 show ?thesis proof (cases "j = 3 + jj") case True then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (0, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by (smt (verit, ccfv_threshold)) also have "... = tps' ! j" using 4 assms(1,2,6,7) 4 True act_onesie tpsL_mvs by simp finally show ?thesis . next case False then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by auto also have "... = tps' ! j" using 4 assms(2,7) False act_Stay len rs_def that tpsL_mvs' - by (smt (z3) add.commute le_add_diff_inverse2) + by (smt (verit) add.commute le_add_diff_inverse2) finally show ?thesis . qed next case 5 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL67_3: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "exec t <#> jj > 0" and "tt = TT - exec t <#> jj" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) (fmt n - Suc tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL67 jj) (0, tps) (Suc tt) = (if fmt n - tt = 0 then 1 else 0, tps')" proof - let ?i = "fmt n - tt" let ?xs = "xs[jj:=(1, None)]" let ?tps = "tpsL t ?xs ?i 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j)" have 1: "Suc ?i = exec t <#> jj" using assms exec_pos_less_TT - by (smt (z3) Suc_diff_le diff_diff_cancel diff_is_0_eq nat_less_le neq0_conv not_less_eq zero_less_diff) + by (smt (verit) Suc_diff_le diff_diff_cancel diff_is_0_eq nat_less_le neq0_conv not_less_eq zero_less_diff) have 2: "?xs ! jj = (1, None)" by (simp add: assms(1) assms(3)) have 3: "length ?xs = k" by (simp add: assms(3)) have "execute (tmL67 jj) (0, tps) (Suc tt) = exe (tmL67 jj) (execute (tmL67 jj) (0, tps) tt)" by simp also have "... = exe (tmL67 jj) (0, ?tps)" using assms execute_tmL67_2a by simp also have "... = sem (cmdL7 jj) (0, ?tps)" using tmL67_def exe_lt_length by simp also have "... = (if fmt n - tt = 0 then 1 else 0, tps')" using sem_cmdL7_3[OF assms(1) _ 3 2 1 assms(5)] assms(8) by simp finally show ?thesis by simp qed lemma sem_cmdL7_4: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 1)" and "Suc i < exec t <#> jj" and "sim_move t ! jj = 0" and "tps' = tpsL t xs (i - 1) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL7 jj) (0, tps) = (if i = 0 then 1 else 0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL7 jj)" using turing_command_cmdL7[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(7) by simp define rs where "rs = read tps" show "fst (cmdL7 jj (read tps)) = (if i = 0 then 1 else 0)" proof (cases "i = 0") case True then have "is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms(2) enc_nth_def read_tpsL_1_bounds rs_def by simp then show ?thesis using True cmdL7_def rs_def by simp next case False then have "\ is_beginning (rs ! 1)" using read_tpsL_1_nth_2k1 assms enc_nth_def exec_pos_less_TT read_tpsL_1 rs_def by (metis (no_types, lifting) less_2_cases_iff nat_1_add_1 not_less_eq plus_1_eq_Suc) then show ?thesis using False cmdL7_def rs_def by simp qed have "i < exec t <#> jj" using assms(5) by simp then have "i < TT" using assms(1) exec_pos_less_TT by (meson Suc_lessD less_trans_Suc) have "rs ! (3 + jj) = \" using rs_def read_tpsL_3 assms by simp moreover have "enc_nth (rs ! 1) (k + jj) = 0" using assms rs_def read_tpsL_1_nth_less_2k[OF `i < TT`, of "k + jj"] sim_move by simp ultimately have "condition7c rs jj" by simp then have *: "snd (cmdL7 jj rs) = [(rs ! 0, Stay), (rs ! 1, Left), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,7) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (rs ! 1, Left) (tps ! 1)" using * rs_def by simp also have "... = tps' ! 1" using assms rs_def tpsL_1 "2" act_Left fst_conv len that tpsL_pos_1 by metis finally show ?thesis using 2 by simp next case 3 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (cmdL7 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,7) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by auto also have "... = tps' ! j" - using 4 assms(2,7) act_Stay len rs_def that tpsL_mvs' by (smt (z3) add.commute le_add_diff_inverse2) + using 4 assms(2,7) act_Stay len rs_def that tpsL_mvs' by (smt (verit) add.commute le_add_diff_inverse2) finally show ?thesis . next case 5 then have "act (cmdL7 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL67_4: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "exec t <#> jj > 0" and "tt \ Suc (Suc (TT - exec t <#> jj))" and "tt \ TT" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) (fmt n - tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL67 jj) (0, tps) tt = (if TT - tt = 0 then 1 else 0, tps')" using assms(7,8,9) proof (induction tt arbitrary: tps' rule: nat_induct_at_least) case base let ?tt = "TT - exec t <#> jj" let ?xs = "xs[jj:=(1, Some 1)]" let ?tps = "tpsL t ?xs (fmt n - Suc ?tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" have "execute (tmL67 jj) (0, tps) (Suc (Suc ?tt)) = exe (tmL67 jj) (execute (tmL67 jj) (0, tps) (Suc ?tt))" by simp also have "... = exe (tmL67 jj) (if fmt n - ?tt = 0 then 1 else 0, ?tps)" using execute_tmL67_3 assms by simp also have "... = sem (cmdL7 jj) (0, ?tps)" using tmL67_def base exe_lt_length by simp finally show ?case using sem_cmdL7_4 assms base.prems(2) by simp next case (Suc tt) let ?tt = "TT - exec t <#> jj" let ?xs = "xs[jj:=(1, Some 1)]" let ?tps = "tpsL t ?xs (fmt n - tt) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" have "execute (tmL67 jj) (0, tps) (Suc tt) = exe (tmL67 jj) (execute (tmL67 jj) (0, tps) tt)" by simp also have "... = exe (tmL67 jj) (if Suc (fmt n) - tt = 0 then 1 else 0, ?tps)" using Suc by simp also have "... = sem (cmdL7 jj) (0, ?tps)" using Suc.prems(1) exe_lt_length tmL67_def by auto finally show ?case using assms sem_cmdL7_4 Suc by simp qed lemma tmL67_left_gt_0: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "exec t <#> jj > 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL67 jj) tps esL67 tps'" proof show 1: "execute (tmL67 jj) (0, tps) (length esL67) = (length (tmL67 jj), tps')" proof (cases "exec t <#> jj = 1") case True then show ?thesis using assms(7) execute_tmL67_3[OF assms(1-6)] esL67_def tmL67_def by simp next case False then have "TT \ Suc (Suc (TT - exec t <#> jj))" using assms(1,6,7) exec_pos_less_TT by (metis Suc_leI add_diff_cancel_right' diff_diff_cancel diff_less nat_less_le plus_1_eq_Suc zero_less_Suc) then show ?thesis using assms(7) execute_tmL67_4[OF assms(1-6), where ?tt=TT] esL67_def tmL67_def by simp qed show "fst (execute (tmL67 jj) (0, tps) tt) < length (tmL67 jj)" if "tt < length esL67" for tt proof - have "tt < TT" using that esL67_def by simp then consider "tt < TT - exec t <#> jj" | "tt = TT - exec t <#> jj" | "tt = Suc (TT - exec t <#> jj)" | "tt \ Suc (Suc (TT - exec t <#> jj))" by linarith then show ?thesis proof (cases) case 1 then show ?thesis using assms execute_tmL67_1 tmL67_def by simp next case 2 then show ?thesis using assms execute_tmL67_2a tmL67_def by simp next case 3 then show ?thesis using assms execute_tmL67_3 tmL67_def `tt < TT` by simp next case 4 then show ?thesis using assms execute_tmL67_4 tmL67_def `tt < TT` by simp qed qed show "execute (tmL67 jj) (0, tps) (Suc tt) <#> 0 = fst (esL67 ! tt) \ execute (tmL67 jj) (0, tps) (Suc tt) <#> 1 = snd (esL67 ! tt)" if "tt < length esL67" for tt proof - have *: "Suc tt \ TT" using that esL67_def by simp then consider "Suc tt < TT - exec t <#> jj" | "Suc tt = TT - exec t <#> jj" | "Suc tt = Suc (TT - exec t <#> jj)" | "Suc tt \ Suc (Suc (TT - exec t <#> jj))" using esL67_def `tt < length esL67` by linarith then show ?thesis proof (cases) case 1 then show ?thesis using execute_tmL67_1[OF assms(1-5), where ?tt="Suc tt"] tmL67_def tpsL_pos_0 tpsL_pos_1 * by simp next case 2 then show ?thesis using assms execute_tmL67_2a[OF assms(1-6), where ?tt="Suc tt"] tmL67_def tpsL_pos_0 tpsL_pos_1 * by simp next case 3 then show ?thesis using assms(6) execute_tmL67_3[OF assms(1-6), where ?tt="tt"] tmL67_def tpsL_pos_0 tpsL_pos_1 * by (smt (verit, ccfv_threshold) add.commute diff_Suc_1 diff_diff_left diff_is_0_eq esL67_at_fmtn esL67_at_lt_fmtn nat_less_le plus_1_eq_Suc prod.collapse prod.inject) next case 4 then show ?thesis using assms(7) execute_tmL67_4[OF assms(1-6) _ *] * tmL67_def tpsL_pos_0 tpsL_pos_1 1 esL67_at_fmtn esL67_at_lt_fmtn esL67_def by (smt (verit, best) One_nat_def Suc_diff_Suc add.commute diff_Suc_1 fst_conv le_neq_implies_less length_append length_map length_rev length_upt list.size(3) list.size(4) minus_nat.diff_0 not_less_eq plus_1_eq_Suc snd_conv) qed qed qed lemma tmL67_left: assumes "jj < k" and "tps = tpsL t xs (fmt n) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL67 jj) tps esL67 tps'" using assms tmL67_left_0 tmL67_left_gt_0 by (cases "exec t <#> jj = 0") simp_all definition "esL47 \ esL46 @ esL67" lemma len_esL47: "length esL47 = 2 * TT + 2" using len_esL46 esL47_def esL67_def by simp lemma tmL47_nonleft: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "sim_move t ! jj \ 0" and "tps' = tpsL t (xs[jj:=(1, Some 0)]) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL47 jj) tps esL47 tps'" unfolding tmL47_def esL47_def using assms by (intro traces_sequential[OF tmL46 tmL67_nonleft]) simp_all lemma tmL47_left: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "sim_move t ! jj = 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL47 jj) tps esL47 tps'" unfolding tmL47_def esL47_def using assms by (intro traces_sequential[OF tmL46 tmL67_left[where ?xs="xs[jj:=(1, Some 0)]"]]) simp_all lemma sem_cmdL8_nonright: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "i < TT" and "sim_move t ! jj \ 2" and "tps' = tpsL t xs (Suc i) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL8 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL8 jj)" using turing_command_cmdL8[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(6) by simp define rs where "rs = read tps" then have "rs ! 1 \ \" using assms by (metis not_one_less_zero read_tpsL_1_bounds(1)) then show "fst (cmdL8 jj rs) = 0" unfolding cmdL8_def by simp have "rs ! (3 + jj) = sim_move t ! jj" using rs_def assms(1,2) read_tpsL_3 by simp moreover have "sim_move t ! jj < 3" using sim_move_def assms(1) direction_to_symbol_less sim_move_nth sim_move_nth_else by (metis One_nat_def not_add_less2 not_less_eq numeral_3_eq_3 plus_1_eq_Suc) ultimately have "condition8d rs jj" using assms(5) \rs ! 1 \ \\ by simp then have *: "snd (cmdL8 jj rs) = [(rs ! 0, Stay), (rs ! 1, Right), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then show ?thesis by (metis "*" act_Stay append.simps(2) assms(2) assms(6) len nth_Cons_0 rs_def that tpsL_0) next case 2 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (rs ! 1, Right) (tps ! 1)" using * rs_def by simp also have "... = tps' ! 1" using act_Right len rs_def assms tpsL_1 that tpsL_pos_1 by (metis "2" add.commute fst_conv plus_1_eq_Suc) also have "... = tps' ! j" using 2 by simp finally show ?thesis . next case 3 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,6) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,6) tpsL_mvs' 4 by simp finally show ?thesis . next case 5 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,6) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma sem_cmdL8_TT: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "i = TT" shows "sem (cmdL8 jj) (0, tps) = (1, tps)" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL8 jj)" using turing_command_cmdL8[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps = 2 * k + 3" using assms(2) by simp define rs where "rs = read tps" then have "rs ! 1 = \" using assms read_tpsL_1 by simp then show "fst (cmdL8 jj rs) = 1" unfolding cmdL8_def by simp have "rs ! (3 + jj) = sim_move t ! jj" using rs_def assms(1,2) read_tpsL_3 by simp moreover have "sim_move t ! jj < 3" using sim_move_def assms(1) direction_to_symbol_less sim_move_nth sim_move_nth_else by (metis One_nat_def not_add_less2 not_less_eq numeral_3_eq_3 plus_1_eq_Suc) ultimately have "condition8c rs jj" using \rs ! 1 = \\ by simp then have *: "snd (cmdL8 jj rs) = [(rs ! 0, Stay), (rs ! 1, Stay), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis using "*" act_Stay len rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] threeplus2k_3[where ?a="(rs ! 0, Stay)"] by (cases) simp_all qed qed lemma execute_tmL78_nonright_le_TT: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "sim_move t ! jj \ 2" and "tt \ TT" and "tps' = tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL78 jj) (0, tps) tt = (0, tps')" using assms(5,6) proof (induction tt arbitrary: tps') case 0 then show ?case using assms(1-4) by simp next case (Suc tt) then have "tt < TT" by simp have "execute (tmL78 jj) (0, tps) (Suc tt) = exe (tmL78 jj) (execute (tmL78 jj) (0, tps) tt)" by simp also have "... = exe (tmL78 jj) (0, tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using Suc by simp also have "... = sem (cmdL8 jj) (0, tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using tmL78_def exe_lt_length by simp finally show ?case using sem_cmdL8_nonright[OF assms(1) _ assms(3) `tt < TT` assms(4)] Suc by simp qed lemma execute_tmL78_nonright_eq_Suc_TT: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "sim_move t ! jj \ 2" and "tps' = tpsL t xs TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL78 jj) (0, tps) (Suc TT) = (1, tps')" using assms sem_cmdL8_TT execute_tmL78_nonright_le_TT[where ?tt=TT] exe_lt_length tmL78_def by simp definition "esL78 \ map (\i. (n + 1, Suc i)) ([0..j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "sim_move t ! jj \ 2" and "tps' = tpsL t xs TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL78 jj) tps esL78 tps'" proof have len: "length esL78 = Suc TT" using esL78_def by simp then show 1: "execute (tmL78 jj) (0, tps) (length esL78) = (length (tmL78 jj), tps')" using assms tmL78_def execute_tmL78_nonright_eq_Suc_TT by simp show "\i. i < length esL78 \ fst (execute (tmL78 jj) (0, tps) i) < length (tmL78 jj)" using len assms execute_tmL78_nonright_le_TT tmL78_def by simp show "(execute (tmL78 jj) (0, tps) (Suc i)) <#> 0 = fst (esL78 ! i) \ (execute (tmL78 jj) (0, tps) (Suc i)) <#> 1 = snd (esL78 ! i)" if "i < length esL78" for i proof (cases "i = TT") case True then have "esL78 ! i = (n + 1, TT)" using esL78_def by (simp add: nth_append) then show ?thesis using assms that tpsL_pos_0 tpsL_pos_1 len 1 True by simp next case False then have "i < TT" using that len by simp moreover from this have "esL78 ! i = (n + 1, Suc i)" using esL78_def nth_map_upt_TT by auto ultimately show ?thesis using tpsL_pos_0 tpsL_pos_1 assms(1-4) execute_tmL78_nonright_le_TT by (metis Suc_leI fst_conv snd_conv) qed qed lemma sem_cmdL8_1: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i < exec t <#> jj" and "sim_move t ! jj = 2" and "tps' = tpsL t xs (Suc i) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL8 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL8 jj)" using turing_command_cmdL8[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(7) by simp have "i < TT" using assms exec_pos_less_TT by (meson Suc_less_eq less_SucI less_trans_Suc) define rs where "rs = read tps" then have "rs ! 1 \ \" using assms `i < TT` by (metis not_one_less_zero read_tpsL_1_bounds(1)) then show "fst (cmdL8 jj rs) = 0" unfolding cmdL8_def by simp have "rs ! (3 + jj) = sim_move t ! jj" using rs_def assms(1,2) read_tpsL_3 by simp moreover have "sim_move t ! jj = 2" using sim_move_def assms(1,6) direction_to_symbol_less sim_move_nth sim_move_nth_else by simp moreover have "enc_nth (rs ! 1) (k + jj) = 0" using assms rs_def read_tpsL_1_nth_less_2k[OF `i < TT`, of "k + jj"] by simp ultimately have "condition8d rs jj" using assms \rs ! 1 \ \\ by simp then have *: "snd (cmdL8 jj rs) = [(rs ! 0, Stay), (rs ! 1, Right), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then show ?thesis using "*" act_Stay append.simps(2) assms len nth_Cons_0 rs_def that tpsL_0 by metis next case 2 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (rs ! 1, Right) (tps ! 1)" using * rs_def by simp also have "... = tps' ! 1" using act_Right len rs_def assms tpsL_1 that tpsL_pos_1 by (metis "2" add.commute fst_conv plus_1_eq_Suc) also have "... = tps' ! j" using 2 by simp finally show ?thesis . next case 3 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,7) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_mvs' 4 by simp finally show ?thesis . next case 5 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL78_1: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 2" and "tt \ exec t <#> jj" and "tps' = tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL78 jj) (0, tps) tt = (0, tps')" using assms(6,7) proof (induction tt arbitrary: tps') case 0 then show ?case using assms(1-5) by simp next case (Suc tt) then have "tt < exec t <#> jj" by simp have "execute (tmL78 jj) (0, tps) (Suc tt) = exe (tmL78 jj) (execute (tmL78 jj) (0, tps) tt)" by simp also have "... = exe (tmL78 jj) (0, tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using Suc by simp also have "... = sem (cmdL8 jj) (0, tpsL t xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j))" using exe_lt_length tmL78_def by simp finally show ?case using assms(1-5) sem_cmdL8_1[where ?i=tt] Suc by simp qed lemma sem_cmdL8_2: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "i = exec t <#> jj" and "sim_move t ! jj = 2" and "tps' = tpsL t (xs[jj:=(1, None)]) (Suc i) 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL8 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL8 jj)" using turing_command_cmdL8[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(7) by simp have "i < TT" using assms exec_pos_less_TT by (meson Suc_less_eq less_SucI less_trans_Suc) define rs where "rs = read tps" then have "rs ! 1 \ \" using assms `i < TT` by (metis not_one_less_zero read_tpsL_1_bounds(1)) then show "fst (cmdL8 jj rs) = 0" unfolding cmdL8_def by simp have "rs ! (3 + jj) = 2" using rs_def read_tpsL_3 assms by simp moreover have "enc_nth (rs ! 1) (k + jj) = 1" using assms rs_def read_tpsL_1_nth_less_2k[OF `i < TT`] by simp ultimately have "condition8a rs jj" using `i < TT` assms(2) read_tpsL_1_bounds rs_def by simp then have *: "snd (cmdL8 jj rs) = [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 0, Right), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 3 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,7) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (enc_upd (rs ! 1) (k + jj) 0, Right) (tps ! 1)" using * rs_def by simp also have "... = tps ! 1 |:=| (enc_upd (rs ! 1) (k + jj) 0) |+| 1" using act_Right' 2 len by simp also have "... = tps' ! 1" using assms zip_cont_enc_upd_None rs_def read_tpsL_1 tpsL_1 zip_cont_def by simp finally show ?thesis using 2 by simp next case 3 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,7) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 show ?thesis proof (cases "j = 3 + jj") case True then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (3, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by (smt (verit)) also have "... = tps' ! j" using 4 assms(2,7) True act_onesie tpsL_mvs by simp finally show ?thesis . next case False then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by auto also have "... = tps' ! j" using 4 assms(2,7) False act_Stay len rs_def that tpsL_mvs' - by (smt (z3) add.commute le_add_diff_inverse2) + by (smt (verit) add.commute le_add_diff_inverse2) finally show ?thesis . qed next case 5 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL78_2: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 2" and "tps' = tpsL t (xs[jj:=(1, None)]) (Suc (exec t <#> jj)) 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL78 jj) (0, tps) (Suc (exec t <#> jj)) = (0, tps')" using assms exe_lt_length tmL78_def execute_tmL78_1 sem_cmdL8_2 by simp lemma sem_cmdL8_3: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. if j = jj then 3 else sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, None)" and "i = Suc (exec t <#> jj)" and "sim_move t ! jj = 2" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) (Suc i) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL8 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL8 jj)" using turing_command_cmdL8[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(7) by simp have "i < TT" using assms exec_pos_less_TT sim_move by (metis (no_types, lifting) add_2_eq_Suc' diff_Suc_1) moreover define rs where "rs = read tps" ultimately have "rs ! 1 \ \" by (metis (no_types, lifting) assms(2) not_one_less_zero read_tpsL_1_bounds(1)) then show "fst (cmdL8 jj (read tps)) = 0" using cmdL8_def rs_def by simp have "rs ! (3 + jj) = 3" by (simp add: rs_def assms(1,2) add.commute read_tpsL_3') moreover have "is_code (rs ! 1)" using assms rs_def read_tpsL_1_nth_less_2k `i < TT` read_tpsL_1_bounds by simp ultimately have "condition7b rs jj" using `i < TT` assms(2) read_tpsL_1_bounds rs_def by simp then have *: "snd (cmdL8 jj rs) = [(rs ! 0, Stay), (enc_upd (rs ! 1) (k + jj) 1, Right), (rs ! 2, Stay)] @ (map (\j. (if j = jj then 2 else rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 0) (tps ! 0)" by simp also have "... = act (rs ! 0, Stay) (tps ! 0)" using * rs_def by simp also have "... = tps ! 0" using act_Stay len rs_def by simp also have "... = tps' ! 0" using assms(2,7) tpsL_0 by simp also have "... = tps' ! j" using 1 by simp finally show ?thesis . next case 2 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (enc_upd (rs ! 1) (k + jj) 1, Right) (tps ! 1)" using * rs_def by simp also have "... = tps ! 1 |:=| (enc_upd (rs ! 1) (k + jj) 1) |+| 1" using act_Right' 2 len by simp also have "... = tps' ! 1" using assms zip_cont_enc_upd_Some_Right rs_def read_tpsL_1 tpsL_1 zip_cont_def by simp finally show ?thesis using 2 by simp next case 3 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,7) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 show ?thesis proof (cases "j = 3 + jj") case True then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (2, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 by (smt (verit, ccfv_SIG) diff_add_inverse) also have "... = tps' ! j" using 4 assms(1,2,6,7) 4 True act_onesie tpsL_mvs by simp finally show ?thesis . next case False then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] 4 diff_add_inverse by auto also have "... = tps' ! j" using 4 assms(2,7) False act_Stay len rs_def that tpsL_mvs' - by (smt (z3) add.commute le_add_diff_inverse2) + by (smt (verit) add.commute le_add_diff_inverse2) finally show ?thesis . qed next case 5 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,7) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL78_3: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 2" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) (Suc (Suc (exec t <#> jj))) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL78 jj) (0, tps) (Suc (Suc (exec t <#> jj))) = (0, tps')" using assms exe_lt_length tmL78_def execute_tmL78_2 sem_cmdL8_3 by simp lemma sem_cmdL8_4: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 1)" and "i > Suc (exec t <#> jj)" and "i < TT" and "sim_move t ! jj = 2" and "tps' = tpsL t xs (Suc i) 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "sem (cmdL8 jj) (0, tps) = (0, tps')" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) (cmdL8 jj)" using turing_command_cmdL8[OF assms(1)] turing_commandD(1) by simp show len: "length tps = 2 * k + 3" using assms(2) by simp show "length tps' = 2 * k + 3" using assms(8) by simp define rs where "rs = read tps" then have "rs ! 1 \ \" using assms by (metis not_one_less_zero read_tpsL_1_bounds(1)) then show "fst (cmdL8 jj rs) = 0" unfolding cmdL8_def by simp have "rs ! (3 + jj) = sim_move t ! jj" using rs_def assms read_tpsL_3 by simp moreover have "sim_move t ! jj = 2" using sim_move_def assms(1,7) direction_to_symbol_less sim_move_nth sim_move_nth_else by simp moreover have "enc_nth (rs ! 1) (k + jj) = 0" using assms rs_def read_tpsL_1_nth_less_2k[OF assms(6), of "k + jj"] sim_move by simp ultimately have "condition8d rs jj" using assms \rs ! 1 \ \\ by simp then have *: "snd (cmdL8 jj rs) = [(rs ! 0, Stay), (rs ! 1, Right), (rs ! 2, Stay)] @ (map (\j. (rs ! (j + 3), Stay)) [0..j. (rs ! (3 + k + j), Stay)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then show ?thesis using "*" act_Stay append.simps(2) assms len nth_Cons_0 rs_def that tpsL_0 by metis next case 2 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 1) (tps ! 1)" by simp also have "... = act (rs ! 1, Right) (tps ! 1)" using * rs_def by simp also have "... = tps' ! 1" using act_Right len rs_def assms tpsL_1 that tpsL_pos_1 2 by (metis add.commute fst_conv plus_1_eq_Suc) also have "... = tps' ! j" using 2 by simp finally show ?thesis . next case 3 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (cmdL8 jj (read tps) [!] 2) (tps ! 2)" by simp also have "... = act (rs ! 2, Stay) (tps ! 2)" using * rs_def by simp also have "... = tps ! 2" using act_Stay len rs_def by simp also have "... = tps' ! 2" using assms(2,8) tpsL_2 by simp also have "... = tps' ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,8) tpsL_mvs' 4 by simp finally show ?thesis . next case 5 then have "act (cmdL8 jj (read tps) [!] j) (tps ! j) = act (rs ! j, Stay) (tps ! j)" using * rs_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp also have "... = tps ! j" using len act_Stay rs_def that by simp also have "... = tps' ! j" using assms(2,8) tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed lemma execute_tmL78_4: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 2" and "tt \ Suc (Suc (exec t <#> jj))" and "tt \ TT" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL78 jj) (0, tps) tt = (0, tps')" using assms(6,7,8) proof (induction tt arbitrary: tps' rule: nat_induct_at_least) case base then show ?case using assms(1-5) execute_tmL78_3 by simp next case (Suc tt) then have "tt < TT" by simp let ?xs = "xs[jj:=(1, Some 1)]" let ?tps = "tpsL t ?xs tt 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" have "execute (tmL78 jj) (0, tps) (Suc tt) = exe (tmL78 jj) (execute (tmL78 jj) (0, tps) tt)" by simp also have "... = exe (tmL78 jj) (0, ?tps)" using Suc by simp also have "... = sem (cmdL8 jj) (0, ?tps)" using tmL78_def exe_lt_length by simp then show ?case using sem_cmdL8_4[where ?i=tt] assms Suc by simp qed lemma execute_tmL78_5: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 2" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "execute (tmL78 jj) (0, tps) (Suc TT) = (1, tps')" proof - have *: "TT \ Suc (Suc (exec t <#> jj))" using exec_pos_less_TT sim_move assms(1,5) by (metis Suc_leI add_2_eq_Suc' add_diff_cancel_left' plus_1_eq_Suc) have "execute (tmL78 jj) (0, tps) (Suc TT) = exe (tmL78 jj) (execute (tmL78 jj) (0, tps) TT)" by simp also have "... = exe (tmL78 jj) (0, tps')" using execute_tmL78_4[OF assms(1-5) *] assms(6) by simp also have "... = sem (cmdL8 jj) (0, tps')" using tmL78_def exe_lt_length by simp finally show ?thesis using assms(1,3,6) sem_cmdL8_TT by simp qed lemma tmL78_right: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 2" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL78 jj) tps esL78 tps'" proof have len: "length esL78 = Suc TT" using esL78_def by simp show "execute (tmL78 jj) (0, tps) (length esL78) = (length (tmL78 jj), tps')" using len execute_tmL78_5 assms tmL78_def by simp show "fst (execute (tmL78 jj) (0, tps) tt) < length (tmL78 jj)" if "tt < length esL78" for tt proof - have "tt < Suc TT" using that len by simp then consider "tt \ exec t <#> jj" | "tt = Suc (exec t <#> jj)" | "tt = Suc (Suc (exec t <#> jj))" | "tt > Suc (Suc (exec t <#> jj))" by linarith then show ?thesis proof (cases) case 1 then show ?thesis using assms execute_tmL78_1 tmL78_def by simp next case 2 then show ?thesis using assms execute_tmL78_2 tmL78_def by simp next case 3 then show ?thesis using assms execute_tmL78_3 tmL78_def by simp next case 4 then show ?thesis using assms execute_tmL78_4 tmL78_def `tt < Suc TT` by simp qed qed show "execute (tmL78 jj) (0, tps) (Suc tt) <#> 0 = fst (esL78 ! tt) \ execute (tmL78 jj) (0, tps) (Suc tt) <#> 1 = snd (esL78 ! tt)" if "tt < length esL78" for tt proof - have *: "Suc tt \ Suc TT" using that esL78_def by simp then consider "Suc tt \ exec t <#> jj" | "Suc tt = Suc (exec t <#> jj)" | "Suc tt = Suc (Suc (exec t <#> jj))" | "Suc tt > Suc (Suc (exec t <#> jj)) \ Suc tt \ TT" | "Suc tt = Suc TT" by linarith then show ?thesis proof (cases) case 1 then have "esL78 ! tt = (n + 1, Suc tt)" using nth_map_upt_TT esL78_def by (metis "*" assms(1) exec_pos_less_TT nat_less_le not_less_eq_eq) then show ?thesis using execute_tmL78_1[OF assms(1-5), where ?tt="Suc tt"] tmL78_def tpsL_pos_0 tpsL_pos_1 * 1 by simp next case 2 then show ?thesis using assms execute_tmL78_2[OF assms(1-5)] tmL78_def tpsL_pos_0 tpsL_pos_1 * by (metis (no_types, lifting) esL78_def exec_pos_less_TT fst_conv nat.inject nth_map_upt_TT snd_conv) next case 3 then have "tt < TT" by (metis add_2_eq_Suc' assms(1) assms(5) diff_Suc_1 exec_pos_less_TT sim_move) then have "esL78 ! tt = (n + 1, Suc tt)" using nth_map_upt_TT esL78_def by auto then show ?thesis using assms(6) execute_tmL78_3[OF assms(1-5)] tmL78_def tpsL_pos_0 tpsL_pos_1 * using 3 by simp next case 4 then have **: "Suc tt \ Suc (Suc (exec t <#> jj))" by simp show ?thesis using execute_tmL78_4[OF assms(1-5) **] tmL78_def tpsL_pos_0 tpsL_pos_1 esL78_def by (metis "4" Suc_le_lessD fst_conv nth_map_upt_TT snd_conv) next case 5 then have "esL78 ! tt = (n + 1, TT)" using esL78_def by (simp add: nth_append) then show ?thesis using assms(6) execute_tmL78_5[OF assms(1-5)] tmL78_def tpsL_pos_0 tpsL_pos_1 esL78_def 5 by simp qed qed qed lemma zip_cont_Stay: assumes "jj < k" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 1" shows "zip_cont t xs = zip_cont t (xs[jj:=(1, Some 1)])" proof fix i let ?xs = "xs[jj := (1, Some 1)]" show "zip_cont t xs i = zip_cont t ?xs i" proof (cases "i < TT") case True then show ?thesis proof (rule zip_cont_eqI) show "\j. j < k \ (exec (t + fst (xs ! j)) <:> j) i = (exec (t + fst (?xs ! j)) <:> j) i" using True assms nth_list_update fst_conv by metis show "\j. j < k \ (case snd (xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) = (case snd (?xs ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0)" using assms sim_move by (metis (no_types, lifting) add.commute add.right_neutral add_diff_cancel_right' nth_list_update_eq nth_list_update_neq option.simps(5) plus_1_eq_Suc snd_conv) qed next case False then show ?thesis by (simp add: zip_cont_def) qed qed lemma tpsL_Stay: assumes "jj < k" and "tps = tpsL t xs i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (1, Some 0)" and "sim_move t ! jj = 1" shows "tps = tpsL t (xs[jj:=(1, Some 1)]) i 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" proof - let ?lhs = "((\zs\, n + 1) # (zip_cont t xs, i) # \fst (exec (t + 1))\ # map (onesie \ (!) (sim_move t)) [0.. (!) (sim_write t)) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < length ?lhs` by fastforce then show ?thesis using zip_cont_Stay assms by (cases) auto qed qed then show ?thesis using assms tpsL_def by simp qed definition "esL48 \ esL47 @ esL78" lemma len_esL48: "length esL48 = 3 * TT + 3" using len_esL47 esL48_def esL78_def by simp lemma tmL48_left: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "sim_move t ! jj = 0" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL48 jj) tps esL48 tps'" unfolding tmL48_def esL48_def using assms by (intro traces_sequential[OF tmL47_left tmL78_nonright[where ?xs="xs[jj:=(1, Some 1)]"]]) simp_all lemma tmL48_right: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "sim_move t ! jj = 2" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL48 jj) tps esL48 tps'" unfolding tmL48_def esL48_def using assms by (intro traces_sequential[OF tmL47_nonleft tmL78_right[where ?xs="xs[jj:=(1, Some 0)]"]]) simp_all lemma tmL48_stay: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "sim_move t ! jj = 1" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL48 jj) tps esL48 tps'" proof - let ?xs = "xs[jj:=(1, Some 0)]" let ?tps = "tpsL t ?xs TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" have "traces (tmL48 jj) tps esL48 ?tps" unfolding tmL48_def esL48_def using assms by (intro traces_sequential[OF tmL47_nonleft tmL78_nonright[where ?xs="?xs"]]) simp_all then show ?thesis using tpsL_Stay[where ?xs="?xs"] assms by simp qed lemma tmL48: assumes "jj < k" and "tps = tpsL t xs 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL48 jj) tps esL48 tps'" proof - consider "sim_move t ! jj = 0" | "sim_move t ! jj = 1" | "sim_move t ! jj = 2" using direction_to_symbol_less sim_move_def assms(1) sim_move_nth sim_move_nth_else by (metis (no_types, lifting) One_nat_def Suc_1 less_Suc_eq less_nat_zero_code numeral_3_eq_3) then show ?thesis using assms tmL48_left tmL48_right tmL48_stay by (cases) simp_all qed definition "esL49 \ esL48 @ map (\i. (n + 1, i)) (rev [0..j. sim_move t ! j) (\j. sim_write t ! j)" and "length xs = k" and "xs ! jj = (0, Some 0)" and "tps' = tpsL t (xs[jj:=(1, Some 1)]) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL49 jj) tps esL49 tps'" unfolding tmL49_def esL49_def proof (rule traces_sequential) let ?tps = "tpsL t (xs[jj:=(1, Some 1)]) TT 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" show "traces (tmL48 jj) tps esL48 ?tps" using assms tmL48 by simp show "traces tm_left_until1 ?tps (map (Pair (n + 1)) (rev [0.. 1 < y \ dec y ! (2 * k + 1) = 1} (?tps ! 1)" using tpsL_1 begin_tape_zip_cont by simp show "map (Pair (n + 1)) (rev [0.. (nat \ nat option) list" where "xs49 j \ replicate j (1, Some 1) @ replicate (k - j) (0, Some 0)" lemma length_xs49: "j \ k \ length (xs49 j) = k" using xs49_def by simp lemma xs49_less: assumes "j \ k" and "i < j" shows "xs49 j ! i = (1, Some 1)" unfolding xs49_def using assms by (simp add: nth_append) lemma xs49_ge: assumes "j \ k" and "i \ j" and "i < k" shows "xs49 j ! i = (0, Some 0)" unfolding xs49_def using assms by (simp add: nth_append) lemma xs49_upd: assumes "j < k" shows "xs49 (Suc j) = (xs49 j)[j := (1, Some 1)]" (is "?lhs = ?rhs") proof (rule nth_equalityI) show "length ?lhs = length ?rhs" by (simp add: Suc_leI assms length_xs49 less_imp_le_nat) show "\i. i < length ?lhs \ ?lhs ! i = ?rhs ! i" using length_xs49 assms xs49_ge xs49_less by (metis less_Suc_eq less_or_eq_imp_le not_less nth_list_update) qed lemma tmL49_upt: assumes "j \ k" and "tps' = tpsL t (xs49 j) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL49_upt j) (tpsL4 t) (concat (replicate j esL49)) tps'" using assms proof (induction j arbitrary: tps') case 0 then show ?case using xs49_def tpsL4_def assms by auto next case (Suc j) then have "j < k" by simp let ?tps = "tpsL t (xs49 j) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" have "tmL49_upt (Suc j) = tmL49_upt j ;; tmL49 j" by simp moreover have "concat (replicate (Suc j) esL49) = concat (replicate j esL49) @ esL49" - by (smt (z3) append.assoc append_replicate_commute append_same_eq concat.simps(2) concat_append + by (smt (verit) append.assoc append_replicate_commute append_same_eq concat.simps(2) concat_append replicate.simps(2)) moreover have "traces (tmL49_upt j ;; tmL49 j) (tpsL4 t) (concat (replicate j esL49) @ esL49) tps'" proof (rule traces_sequential) show "traces (tmL49_upt j) (tpsL4 t) (concat (replicate j esL49)) ?tps" using Suc by simp show "traces (tmL49 j) ?tps esL49 tps'" using tmL49[OF `j < k`, where ?tps="?tps"] length_xs49 xs49_upd Suc \j < k\ xs49_ge by simp qed ultimately show ?case by simp qed definition "esL49_upt \ concat (replicate k esL49)" lemma length_concat_replicate: "length (concat (replicate m xs)) = m * length xs" by (induction m) simp_all lemma len_esL49_upt: "length esL49_upt = k * (4 * TT + 4)" using len_esL49 esL49_upt_def length_concat_replicate[of k esL49] by simp corollary tmL49_upt': assumes "tps' = tpsL t (xs49 k) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" shows "traces (tmL49_upt k) (tpsL4 t) esL49_upt tps'" using tmL49_upt[of k] assms esL49_upt_def by simp definition "esL9 t \ esL4 t @ esL49_upt" lemma len_esL9: "length (esL9 t) = k * (4 * TT + 4) + t + 2 * TT + 4" using len_esL4 len_esL49_upt esL9_def by simp lemma xs49_k: "xs49 k = replicate k (1, Some 1)" using xs49_def by simp definition "tpsL9 t \ tpsL t (replicate k (1, Some 1)) 0 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" lemma tmL9: assumes "t < TT" shows "traces tmL9 (tpsL0 t) (esL9 t) (tpsL9 t)" unfolding tmL9_def esL9_def using assms tmL4 tmL49_upt' by (intro traces_sequential) (auto simp add: tpsL9_def xs49_k) definition "esL10 t \ esL9 t @ esC t" lemma len_esL10: "length (esL10 t) = k * (4 * TT + 4) + 2 * t + 2 * TT + 5" using len_esL9 len_esL9 esL10_def esC_def by simp definition "tpsL10 t \ tpsL t (replicate k (1, Some 1)) t 1 (\j. sim_move t ! j) (\j. sim_write t ! j)" lemma tmL10: assumes "t < TT" shows "traces tmL10 (tpsL0 t) (esL10 t) (tpsL10 t)" unfolding tmL10_def esL10_def proof (rule traces_sequential[OF tmL9[OF assms]]) have "t \ TT" using assms by simp then show "traces tmC (tpsL9 t) (esC t) (tpsL10 t)" using tmC_general tpsL9_def tpsL10_def by simp qed definition "tpsL11 t \ tpsL (Suc t) (replicate k (0, Some 0)) t 0 (\j. sim_move t ! j) (\j. sim_write t ! j)" lemma enc_upd_2k: assumes "dec n = (map f [0..j. (exec (t + fst (xs1 ! j)) <:> j) t) [0..j. case snd (xs1 ! j) of None \ 0 | Some d \ if t = exec (t + d) <#> j then 1 else 0) [0..j. (exec (t + 1) <:> j) t) [0..j. case Some 1 of None \ 0 | Some d \ if t = exec (t + d) <#> j then 1 else 0) [0..j. (exec (Suc t) <:> j) t) [0..j. if t = exec (Suc t) <#> j then 1 else 0) [0..j. (exec (Suc t + fst (xs0 ! j)) <:> j) t) [0..j. case snd (xs0 ! j) of None \ 0 | Some d \ if t = exec (Suc t + d) <#> j then 1 else 0) [0..j. (exec (Suc t) <:> j) t) [0..j. case Some 0 of None \ 0 | Some d \ if t = exec (Suc t + d) <#> j then 1 else 0) [0..j. (exec (Suc t) <:> j) t) [0..j. if t = exec (Suc t) <#> j then 1 else 0) [0.. t \ i < TT" | "i \ TT" using assms(1) by linarith then show "((zip_cont t xs1)(t := enc_upd (zip_cont t xs1 t) (2 * k) 1)) i = zip_cont (Suc t) xs0 i" proof (cases) case 1 then show ?thesis using enc_upd_zip_cont assms by simp next case 2 then have "i \ t" "i < TT" by simp_all have xs1: "fst (xs1 ! j) = 1" "snd (xs1 ! j) = Some 1" if "j < k" for j using assms(2) that by simp_all have "zip_cont t xs1 i = enc (map (\j. (exec (t + fst (xs1 ! j)) <:> j) i) [0..j. case snd (xs1 ! j) of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..j. (exec (t + 1) <:> j) i) [0..j. case Some 1 of None \ 0 | Some d \ if i = exec (t + d) <#> j then 1 else 0) [0..j. (exec (Suc t) <:> j) i) [0..j. if i = exec (Suc t) <#> j then 1 else 0) [0..j. (exec (Suc t + fst (xs0 ! j)) <:> j) i) [0..j. case snd (xs0 ! j) of None \ 0 | Some d \ if i = exec (Suc t + d) <#> j then 1 else 0) [0..j. (exec (Suc t) <:> j) i) [0..j. case Some 0 of None \ 0 | Some d \ if i = exec (Suc t + d) <#> j then 1 else 0) [0..j. (exec (Suc t) <:> j) i) [0..j. if i = exec (Suc t) <#> j then 1 else 0) [0.. t` by simp then show ?thesis using 1 `i \ t` by simp next case 3 then show ?thesis using zip_cont_def assms(1) by simp qed qed lemma sem_cmdL11: assumes "t < TT" shows "sem cmdL11 (0, tpsL10 t) = (1, tpsL11 t)" proof (rule semI[of "2 * k + 3"]) show "proper_command (2 * k + 3) cmdL11" using cmdL11_def by simp show len: "length (tpsL10 t) = 2 * k + 3" "length (tpsL11 t) = 2 * k + 3" using tpsL10_def tpsL11_def by simp_all show "fst (cmdL11 (read (tpsL10 t))) = 1" using cmdL11_def by simp let ?tps = "tpsL10 t" let ?xs = "replicate k (1::nat, Some 1::nat option)" have tps1: "?tps ! 1 = (zip_cont t ?xs, t)" using tpsL_1 tpsL10_def by simp have tps1': "tpsL11 t ! 1 = (zip_cont (Suc t) (replicate k (0, Some 0)), t)" using tpsL_1 tpsL11_def by simp let ?rs = "read ?tps" have "is_code (?rs ! 1)" using tpsL10_def assms read_tpsL_1_bounds by simp have rs1: "?rs ! 1 = zip_cont t ?xs t" using tps1 read_def tpsL_def tpsL10_def by force show " act (cmdL11 ?rs [!] j) (?tps ! j) = tpsL11 t ! j" if "j < 2 * k + 3" for j proof - consider "j = 0" | "j = 1" | "j = 2" | "3 \ j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < 2 * k + 3` by linarith then show ?thesis proof (cases) case 1 then show ?thesis using tpsL10_def tpsL11_def read_tpsL_0 cmdL11_def act_Stay len(1) that tpsL_0 by (smt (verit) append_Cons nth_Cons_0 prod.sel(2)) next case 2 then have "act (cmdL11 ?rs [!] j) (?tps ! j) = act (cmdL11 ?rs [!] 1) (?tps ! 1)" by simp also have "... = act (enc_upd (?rs ! 1) (2 * k) 1, Stay) (?tps ! 1)" using cmdL11_def `is_code (?rs ! 1)` by simp also have "... = (?tps ! 1) |:=| enc_upd (?rs ! 1) (2 * k) 1" by (simp add: act_Stay tps1 "2" act_Stay' len(1) that) also have "... = tpsL11 t ! 1" using enc_upd_zip_cont_upd rs1 tps1' tps1 assms by simp finally show ?thesis using 2 by simp next case 3 then have "act (cmdL11 ?rs [!] j) ((?tps) ! j) = act (cmdL11 ?rs [!] 2) (?tps ! 2)" by simp also have "... = act (?rs ! 2, Stay) (?tps ! 2)" using cmdL11_def by simp also have "... = ?tps ! 2" using act_Stay len by simp also have "... = (tpsL11 t) ! 2" using tpsL_2 tpsL11_def tpsL10_def by simp also have "... = (tpsL11 t) ! j" using 3 by simp finally show ?thesis . next case 4 then have "act (cmdL11 ?rs [!] j) (?tps ! j) = act (?rs ! j, Stay) (?tps ! j)" using cmdL11_def threeplus2k_2[where ?a="(?rs ! 0, Stay)"] by simp also have "... = (tpsL10 t) ! j" using len act_Stay that by simp also have "... = (tpsL11 t) ! j" using tpsL11_def tpsL10_def tpsL_mvs' 4 by simp finally show ?thesis . next case 5 then have "act (cmdL11 ?rs [!] j) (?tps ! j) = act (?rs ! j, Stay) (?tps ! j)" using cmdL11_def threeplus2k_3[where ?a="(?rs ! 0, Stay)"] by simp also have "... = (tpsL10 t) ! j" using len act_Stay that by simp also have "... = (tpsL11 t) ! j" using tpsL11_def tpsL10_def tpsL_symbs' 5 by simp finally show ?thesis . qed qed qed definition "esL11 t \ esL10 t @ [(n + 1, t)]" lemma len_esL11: "length (esL11 t) = k * (4 * TT + 4) + 2 * t + 2 * TT + 6" using len_esL10 esL11_def by simp lemma tmL11: assumes "t < TT" shows "traces tmL11 (tpsL0 t) (esL11 t) (tpsL11 t)" unfolding tmL11_def esL11_def proof (rule traces_sequential[OF tmL10[OF assms]]) let ?cmd = "[cmdL11]" let ?es = "[(n + 1, t)]" show "traces ?cmd (tpsL10 t) ?es (tpsL11 t)" proof show 1: "execute ?cmd (0, tpsL10 t) (length ?es) = (length ?cmd, tpsL11 t)" proof - have "execute ?cmd (0, tpsL10 t) (length ?es) = exe ?cmd (0, tpsL10 t)" by simp also have "... = sem cmdL11 (0, tpsL10 t)" using exe_lt_length cmdL11_def by simp finally show ?thesis using sem_cmdL11[OF assms] by simp qed show "\i. i < length [(n + 1, t)] \ fst (execute ?cmd (0, tpsL10 t) i) < length ?cmd" by simp show "\i. i < length [(n + 1, t)] \ execute ?cmd (0, tpsL10 t) (Suc i) <#> 0 = fst (?es ! i) \ execute ?cmd (0, tpsL10 t) (Suc i) <#> 1 = snd (?es ! i)" using 1 tpsL11_def tpsL_pos_0 tpsL_pos_1 by (metis One_nat_def add.commute fst_conv less_Suc0 list.size(3) list.size(4) nth_Cons_0 plus_1_eq_Suc snd_conv) qed qed definition "esL12 t \ esL11 t @ map (\i. (n + 1, i)) (rev [0.. tpsL (Suc t) (replicate k (0, Some 0)) 0 0 (\j. sim_move t ! j) (\j. sim_write t ! j)" lemma tmL12: assumes "t < TT" shows "traces tmL12 (tpsL0 t) (esL12 t) (tpsL12 t)" unfolding tmL12_def esL12_def proof (rule traces_sequential[OF tmL11[OF assms]]) show "traces tm_left_until1 (tpsL11 t) (map (Pair (n + 1)) (rev [0.. 1 < y \ dec y ! (2 * k + 1) = 1} (tpsL11 t ! 1)" using tpsL_1 begin_tape_zip_cont tpsL11_def by simp show "map (Pair (n + 1)) (rev [0.. tpsL (Suc t) (replicate k (0, Some 0)) 0 0 (\j. 0) (\j. 0)" definition "esL13 t \ esL12 t @ [(n + 1, 0)]" lemma len_esL13: "length (esL13 t) = k * (4 * TT + 4) + 3 * t + 2 * TT + 8" using len_esL12 esL13_def by simp lemma tmL13: assumes "t < TT" shows "traces tmL13 (tpsL0 t) (esL13 t) (tpsL13 t)" unfolding tmL13_def esL13_def proof (rule traces_sequential[OF tmL12[OF assms]]) show "traces (tm_write_many {3..<2 * k + 3} 0) (tpsL12 t) [(n + 1, 0)] (tpsL13 t)" proof (rule traces_tm_write_manyI[where ?k="2*k+3"]) show "0 \ {3..<2 * k + 3}" by simp show "\j\{3..<2 * k + 3}. j < 2 * k + 3" by simp show "2 \ 2 * k + 3" by simp show "length (tpsL12 t) = 2 * k + 3" "length (tpsL13 t) = 2 * k + 3" using tpsL12_def tpsL13_def length_tpsL by simp_all show "tpsL13 t ! j = tpsL12 t ! j |:=| 0" if "j \ {3..<2 * k + 3}" for j proof (cases "j < k + 3") case True then have "3 \ j \ j < k + 3" using that by simp then show ?thesis using tpsL13_def tpsL12_def tpsL_mvs' onesie_write by simp next case False then have "k + 3 \ j \ j < 2 * k + 3" using that by simp then show ?thesis using tpsL13_def tpsL12_def tpsL_symbs' onesie_write by simp qed show "tpsL13 t ! j = tpsL12 t ! j" if "j < 2 * k + 3" "j \ {3..<2 * k + 3}" for j proof - from that have "j < 3" by simp then show ?thesis using tpsL13_def tpsL12_def tpsL_def less_Suc_eq numeral_3_eq_3 by auto qed show "[(n + 1, 0)] = [(tpsL12 t :#: 0, tpsL12 t :#: 1)]" using tpsL12_def tpsL_pos_0 tpsL_pos_1 by simp qed qed corollary tmL13': assumes "t < TT" shows "traces tmL13 (tpsC1 t) (esL13 t) (tpsL13 t)" using tmL13 tpsC1_def tpsL0_def assms by simp definition "esLoop_while t \ esC t @ [(tpsC1 t :#: 0, tpsC1 t :#: 1)] @ esL13 t @ [(tpsL13 t :#: 0, tpsL13 t :#: 1)]" definition "esLoop_break \ (esC TT) @ [(tpsC1 TT :#: 0, tpsC1 TT :#: 1)]" lemma len_esLoop_while: "length (esLoop_while t) = k * (4 * TT + 4) + 4 * t + 2 * TT + 11" using len_esL13 esC_def esLoop_while_def by simp lemma tmLoop_while: assumes "t < TT" shows "trace tmLoop (0, tpsC0 t) (esLoop_while t) (0, tpsL13 t)" unfolding tmLoop_def proof (rule tm_loop_sem_true_tracesI[OF tmC tmL13']) show "t \ TT" and "t < TT" using assms by simp_all show "0 < read (tpsC1 t) ! 1" using tpsC1_def read_tpsL_1_bounds(1) assms by (metis gr0I not_one_less_zero) show "esLoop_while t = esC t @ [(tpsC1 t :#: 0, tpsC1 t :#: 1)] @ esL13 t @ [(tpsL13 t :#: 0, tpsL13 t :#: 1)]" using esLoop_while_def by simp qed lemma tmLoop_while_end: "trace tmLoop (0, tpsC0 0) (concat (map esLoop_while [0.. TT" by simp show "\ 0 < read (tpsC1 TT) ! 1" using read_tpsL_1 tpsC1_def by simp show "esLoop_break = esC (TT) @ [(tpsC1 TT :#: 0, tpsC1 TT :#: 1)]" using esLoop_break_def by simp qed definition "esLoop \ concat (map esLoop_while [0.. TT \ length (concat (map esLoop_while [0.. u * (k * (4 * TT + 4) + 4 * TT + 2 * TT + 11)" using len_esLoop_while by (induction u) simp_all lemma len_esLoop2: "length (concat (map esLoop_while [0.. TT * (k * (4 * TT + 4) + 4 * TT + 2 * TT + 11)" using len_esLoop1[of TT] by simp lemma len_esLoop3: "length esLoop \ TT * (k * (4 * TT + 4) + 4 * TT + 2 * TT + 11) + TT + 2" using len_esLoop2 esC_def esLoop_def esLoop_break_def by simp lemma len_esLoop: "length esLoop \ 28 * k * TT * TT" proof - have "length esLoop \ TT * (k * (4 * TT + 4) + 4 * TT + 2 * TT + 11) + TT + 2" using len_esLoop3 . also have "... = TT * (k * (4 * TT + 4) + 6 * TT + 11) + TT + 2" by simp also have "... \ TT * (k * (4 * TT + 4) + 6 * TT + 11) + 3 * TT" by simp also have "... = TT * k * (4 * TT + 4) + TT * 6 * TT + TT * 11 + 3 * TT" by algebra also have "... = TT * k * (4 * TT + 4) + TT * 6 * TT + 14 * TT" by simp also have "... = k * 4 * TT * TT + k * 4 * TT + 6 * TT * TT + 14 * TT" by algebra also have "... \ k * 4 * TT * TT + k * 4 * TT * TT + 6 * TT * TT + 14 * TT * TT" by simp also have "... = (k * 4 + k * 4 + 6 + 14) * TT * TT" by algebra also have "... = (k * 8 + 20) * TT * TT" by algebra also have "... \ 28 * k * TT * TT" proof - have "k * 8 + 20 \ 28 * k" using k_ge_2 by simp then show ?thesis by (meson mult_le_mono1) qed finally show ?thesis by simp qed lemma tmLoop: "traces tmLoop (tpsC0 0) esLoop (tpsC1 TT)" unfolding esLoop_def using traces_additive[OF tmLoop_while_end tmLoop_break] . lemma tps9_tpsC0: "tps9 = tpsC0 0" using tps9_def tpsC0_def tps9_tpsL by simp definition "es10 \ es9 @ esLoop" lemma len_es10: "length es10 \ length (es_fmt n) + 40 * k * TT * TT" proof - have "length es10 \ length (es_fmt n) + 2 * TT + 2 * n + 8 + 28 * k * TT * TT" using len_es9 len_esLoop es10_def by simp also have "... \ length (es_fmt n) + 2 * TT + 2 * TT + 8 + 28 * k * TT * TT" proof - have "2 * n \ 2 * TT" using fmt_ge_n Suc_mult_le_cancel1 le_SucI numeral_2_eq_2 by metis then show ?thesis by simp qed also have "... \ length (es_fmt n) + 12 * TT * TT + 28 * k * TT * TT" by simp also have "... \ length (es_fmt n) + 12 * k * TT * TT + 28 * k * TT * TT" proof - have "12 \ 12 * k" using k_ge_2 by simp then have "12 * TT * TT \ 12 * k * TT * TT" using mult_le_mono1 by presburger then show ?thesis by simp qed also have "... = length (es_fmt n) + 40 * k * TT * TT" by linarith finally show ?thesis . qed lemma tm10: "traces tm10 tps0 es10 (tpsC1 TT)" unfolding tm10_def es10_def using traces_sequential[OF tm9] tps9_tpsC0 tmLoop by simp subsubsection \Cleaning up the output\ abbreviation "tps10 \ tpsC1 TT" definition "es11 \ es10 @ map (\i. (n + 1, i)) (rev [0.. length (es_fmt n) + 40 * k * TT * TT + Suc TT" using es11_def len_es10 by simp definition "tps11 \ tps10[1 := ltransplant (tps10 ! 1) (tps10 ! 1) ec1 TT]" lemma tm11: "traces tm11 tps0 es11 tps11" unfolding tm11_def es11_def proof (rule traces_sequential[OF tm10]) show "traces (tm_ltrans_until 1 1 StartSym ec1) (tpsC1 (Suc (fmt n))) (map (Pair (n + 1)) (rev [0..h tpsC1 (Suc (fmt n)) :#: 1" using tpsC1_def tpsL_def by simp show "map (Pair (n + 1)) (rev [0..i. (tps10 :#: 0, tps10 :#: 1 - Suc i)) [0..i. (tps10 :#: 0, tps10 :#: 1 - Suc i)) [0..i. (n + 1, TT - Suc i)) [0..i. (c1, c2 - Suc i)) [0..i. (tps10 :#: 0, tps10 :#: 1 - Suc i)) [0.. es11 @ [(n + 1, 1)]" text \ The upper bound on the length of the trace will help us establish an upper bound of the total running time. \ lemma length_es12: "length es12 \ length (es_fmt n) + 43 * k * TT * TT" proof - have "length es12 \ length (es_fmt n) + 40 * k * TT * TT + 3 * TT * TT" using es12_def len_es11 by simp moreover have "3 * TT * TT \ 3 * k * TT * TT" proof - have "3 \ 3 * k" using k_ge_2 by simp then show ?thesis by (meson mult_le_mono1) qed ultimately show ?thesis by linarith qed definition "tps12 \ tps11[1 := tps11 ! 1 |:=| (ec1 (tps11 :.: 1)) |+| 1]" lemma tm12: "traces tm12 tps0 es12 tps12" unfolding tm12_def es12_def proof (rule traces_sequential[OF tm11]) show "traces (tm_rtrans 1 ec1) tps11 [(n + 1, 1)] tps12" proof (rule traces_tm_rtrans_1I) show "1 < length tps11" using tps11_def tpsC1_def by simp show "[(n + 1, 1)] = [(tps11 :#: 0, tps11 :#: 1 + 1)]" using tps11_def tpsC1_def tpsL_pos_0 tpsL_pos_1 ltransplant_def by simp show "tps12 = tps11[1 := tps11 ! 1 |:=| ec1 (tps11 :.: 1) |+| 1]" using tps12_def by simp qed qed lemma tps11_0: "(tps11 ::: 1) 0 = (zip_cont TT (replicate k (0, Some 0))) 0" using tps11_def tpsC1_def tpsL_def ltransplant_def by simp lemma tps11_gr0_exec: assumes "i > 0" shows "(tps11 ::: 1) i = (exec TT <:> 1) i" proof - let ?tp = "tps10 ! 1" let ?xs = "replicate k (0, Some 0)" have 1: "tps11 ! 1 = ltransplant ?tp ?tp ec1 TT" using tps11_def tpsC1_def tpsL_def by simp have 2: "tps10 :#: 1 = TT" using tpsC1_def tpsL_def by simp show ?thesis proof (cases "i \ TT") case le_TT: True then have "0 < i \ i \ TT" using assms by simp then have *: "(tps11 ::: 1) i = ec1 (fst ?tp i)" using 1 tpsC1_def tpsL_def ltransplant_def by simp show ?thesis proof (cases "i = TT") case True then have "\ is_code ((zip_cont TT ?xs) i)" by (simp add: zip_cont_eq_0) then have "(tps11 ::: 1) i = 0" using * 2 True tpsC1_at_T by simp moreover have "(exec TT <:> 1) TT = 0" using tps_ge_TT_0 by simp ultimately show ?thesis using True by simp next case False then have "i < TT" using le_TT by simp then have "fst ?tp i = (zip_cont TT ?xs) i" using tpsC1_def tpsL_def by simp then have "(tps11 ::: 1) i = ec1 ((zip_cont TT ?xs) i)" using * by simp moreover have "is_code ((zip_cont TT ?xs) i)" using zip_cont_gt_1 zip_cont_less `i < TT` by simp ultimately have "(tps11 ::: 1) i = enc_nth ((zip_cont TT ?xs) i) 1" by simp then have "(tps11 ::: 1) i = (exec TT <:> 1) i" using enc_nth_def dec_zip_cont_less_k[OF `i < TT`] k_ge_2 by simp then show ?thesis by simp qed next case False then have "(tps11 ::: 1) i = 0" using 1 tpsC1_def tpsL_def ltransplant_def zip_cont_eq_0 by force moreover have "(exec TT <:> 1) i = 0" using False tps_ge_TT_0 by simp ultimately show ?thesis by simp qed qed definition "tps12' \ [(\zs\, n + 1), (exec TT <:> 1, 1), \fst (exec TT)\] @ map (\i. \\\) [0..i. \\\) [0.. j \ j < k + 3" | "k + 3 \ j \ j < 2 * k + 3" using `j < length tps12'` by linarith then show ?thesis proof (cases) case 1 then show ?thesis unfolding tps12'_def tps12_def tps11_def tpsC1_def tpsL_def by simp next case 2 then have lhs: "tps12' ! j = (\i. (exec TT <:> 1) i, 1)" using tps12'_def by simp let ?tp = "tps10 ! 1" let ?xs = "replicate k (0, Some 0)" have "tps11 :#: 1 = 0" using tps11_def ltransplant_def tpsC1_def tpsL_pos_1 by simp have rhs: "tps12 ! j = (ltransplant ?tp ?tp ec1 TT) |:=| (ec1 (tps11 :.: 1)) |+| 1" using tps12_def 2 \length tps12' = length tps12\ tps11_def that by simp have tps10: "tps10 ! j = (zip_cont TT ?xs, TT)" using tpsC1_def 2 tpsL_1 by simp show "tps12' ! j = tps12 ! j" proof show "tps12' :#: j = tps12 :#: j" using lhs rhs ltransplant_def tps10 2 by simp have tps12: "tps12 ! 1 = tps11 ! 1 |:=| (ec1 (tps11 :.: 1)) |+| 1" using tps12_def "2" \length tps12' = length tps12\ that by auto have "(tps12' ::: 1) i = (tps12 ::: 1) i" for i proof (cases "i = 0") case True then have "(tps12 ::: 1) i = ec1 (tps11 :.: 1)" using tps12 `tps11 :#: 1 = 0` by simp moreover have "(tps11 :.: 1) = (zip_cont TT ?xs) 0" using tps11_0 `tps11 :#: 1 = 0` by simp ultimately have "(tps12 ::: 1) i = ec1 ((zip_cont TT ?xs) 0)" by simp moreover have "is_code ((zip_cont TT ?xs) 0)" using zip_cont_gt_1 zip_cont_less by simp ultimately have "(tps12 ::: 1) i = enc_nth ((zip_cont TT ?xs) 0) 1" by simp then have "(tps12 ::: 1) i = enc_nth (zip_cont TT ?xs i) 1" using True by simp then have "(tps12 ::: 1) i = (exec TT <:> 1) i" using enc_nth_def dec_zip_cont_less_k True k_ge_2 by simp then show ?thesis using tps12'_def by simp next case False then have "(tps12 ::: 1) i = (tps11 ::: 1) i" using tps12 `tps11 :#: 1 = 0` by simp then have "(tps12 ::: 1) i = (exec TT <:> 1) i" using False tps11_gr0_exec by simp moreover have "(tps12' ::: 1) i = (exec TT <:> 1) i" using tps12'_def by simp ultimately show ?thesis by simp qed then show "tps12' ::: j = tps12 ::: j" using 2 by auto qed next case 3 then show ?thesis unfolding tps12'_def tps12_def tps11_def tpsC1_def tpsL_def by simp next case 4 then show ?thesis unfolding tps12'_def tps12_def tps11_def tpsC1_def using tpsL_mvs' threeplus2k_2[where ?a="(\zs\, n + 1)"] by simp next case 5 then show ?thesis unfolding tps12'_def tps12_def tps11_def tpsC1_def using tpsL_symbs' threeplus2k_3[where ?a="(\zs\, n + 1)"] by simp qed qed qed lemma tm12': "traces tm12 tps0 es12 tps12'" using tm12 tps12' by simp end (* context for zs *) subsection \Shrinking the Turing machine to two tapes\ text \ The simulator TM @{const tm12} has $2k + 3$ tapes, of which $2k + 1$ are immobile and thus can be removed by the memorization-in-states technique, resulting in a two-tape TM. \ lemma immobile_tm12: assumes "j > 1" and "j < 2 * k + 3" shows "immobile tm12 j (2 * k + 3)" proof - have "immobile tm1 j (2 * k + 3)" unfolding tm1_def using immobile_append_tapes[of j "2*k+3", OF _ _ _ fmt(1)] assms by simp moreover have "immobile tm1_2 j (2 * k + 3)" using tm1_2_def tm_const_until_def immobile_tm_trans_until assms by simp ultimately have "immobile tm2 j (2 * k + 3)" using tm2_def immobile_sequential tm1_2_tm tm1_tm by simp then have "immobile tm3 j (2 * k + 3)" using tm3_def immobile_sequential[OF tm2_tm] tm_start_tm immobile_tm_start assms G'_ge by simp then have "immobile tm4 j (2 * k + 3)" using tm4_def immobile_sequential[OF tm3_tm tm3_4_tm] immobile_tm_write assms by simp then have "immobile tm5 j (2 * k + 3)" using tm5_def immobile_sequential[OF tm4_tm] G'_ge(1) immobile_tm_right tm_right_tm assms by simp then have "immobile tm6 j (2 * k + 3)" using tm6_def immobile_sequential[OF tm5_tm tm5_6_tm] immobile_tm_trans_until tm5_6_def assms by simp then have "immobile tm7 j (2 * k + 3)" using tm7_def immobile_sequential[OF tm6_tm tm_left_until1_tm] immobile_tm_left_until assms by simp then have "immobile tm8 j (2 * k + 3)" using tm8_def immobile_sequential[OF tm7_tm] immobile_tm_write assms G'_ge tm_write_tm by simp then have 9: "immobile tm9 j (2 * k + 3)" using tm9_def immobile_sequential[OF tm8_tm] immobile_tm_write_many tm_write_many_tm k_ge_2 G'_ge assms by simp have C: "immobile tmC j (2 * k + 3)" unfolding tmC_def tm_right_until_def tm_cp_until_def using tm_cp_until_tm immobile_tm_trans_until G'_ge(1) assms by simp have "cmdL2 rs [~] j = Stay" if "length rs = 2 * k + 3" for rs proof (cases "rs ! 1 = \") case True then show ?thesis using cmdL2_def assms that by simp next case False then consider "j = 2" | "3 \ j \ j < 3 + k" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis proof (cases) case 1 then show ?thesis using cmdL2_def assms that by simp next case 2 then show ?thesis using assms that cmdL2_at_3 False by simp next case 3 then show ?thesis using assms that cmdL2_at_4 False by simp qed qed then have "immobile tmL1_2 j (2 * k + 3)" using tmL1_2_def by auto then have "immobile tmL2 j (2 * k + 3)" unfolding tmL2_def tmL1_def using tm_left_until1_tm immobile_tm_left_until tmL2_tm immobile_sequential tmL1_2_tm assms by auto moreover have "cmdL3 rs [~] j = Stay" if "length rs = 2 * k + 3" for rs proof - consider "j = 2" | "3 \ j \ j < 3 + k" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis proof (cases) case 1 then show ?thesis using cmdL3_def assms that by simp next case 2 then show ?thesis using assms that cmdL3_at_3a cmdL3_at_3b by (metis (no_types, lifting) add.commute prod.sel(2)) next case 3 then show ?thesis using assms that cmdL3_at_4a cmdL3_at_4b by (metis (no_types, lifting) add.commute prod.sel(2)) qed qed ultimately have "immobile tmL3 j (2 * k + 3)" unfolding tmL3_def using tmL2_tm immobile_sequential assms tmL2_3_def tmL2_3_tm immobile_def by simp then have L4: "immobile tmL4 j (2 * k + 3)" unfolding tmL4_def using tmL3_tm immobile_sequential assms tm_left_until1_tm immobile_tm_left_until by auto have "(cmdL5 jj) rs [~] j = Stay" if "length rs = 2 * k + 3" and "jj < k" for rs jj proof (cases "rs ! 1 = \") case True then show ?thesis using cmdL5_eq_0 assms that by simp next case False then consider "j = 2" | "3 \ j \ j < 3 + k" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis proof (cases) case 1 then show ?thesis using that by (simp add: cmdL5_def) next case 2 then show ?thesis using assms that cmdL5_at_3 False by simp next case 3 then show ?thesis using assms that cmdL5_at_4 False by simp qed qed then have "immobile (tmL45 jj) j (2 * k + 3)" if "jj < k" for jj by (auto simp add: that tmL45_def) then have L46: "immobile (tmL46 jj) j (2 * k + 3)" if "jj < k" for jj using tmL46_def immobile_sequential[OF tmL45_tm] tm_left_tm immobile_tm_left assms that k_ge_2 G'_ge by simp have "(cmdL7 jj) rs [~] j = Stay" if "length rs = 2 * k + 3" and "jj < k" for rs jj proof - consider (a) "condition7a rs jj" | (b) "condition7b rs jj" | (c) "condition7c rs jj" by blast then show ?thesis proof (cases) case a consider "j = 2" | "3 \ j \ j < k + 3" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis using cmdL7_def a threeplus2k_2[of _ _ "(rs ! 0, Stay)"] threeplus2k_3[of _ _ "(rs ! 0, Stay)"] by (cases) simp_all next case b consider "j = 2" | "3 \ j \ j < k + 3" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis using cmdL7_def b threeplus2k_2[of _ _ "(rs ! 0, Stay)"] threeplus2k_3[of _ _ "(rs ! 0, Stay)"] by (cases) simp_all next case c consider "j = 2" | "3 \ j \ j < k + 3" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis using cmdL7_def c threeplus2k_2[of _ _ "(rs ! 0, Stay)"] threeplus2k_3[of _ _ "(rs ! 0, Stay)"] by (cases) simp_all qed qed then have "immobile (tmL67 jj) j (2 * k + 3)" if "jj < k" for jj by (auto simp add: that tmL67_def) then have L47: "immobile (tmL47 jj) j (2 * k + 3)" if "jj < k" for jj using tmL47_def immobile_sequential[OF tmL46_tm tmL67_tm] L46 assms that by simp have "(cmdL8 jj) rs [~] j = Stay" if "length rs = 2 * k + 3" and "jj < k" for rs jj proof - consider (a) "condition8a rs jj" | (b) "condition8b rs jj" | (c) "condition8c rs jj" | (d) "condition8d rs jj" by blast then show ?thesis proof (cases) case a consider "j = 2" | "3 \ j \ j < k + 3" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis using cmdL8_def a threeplus2k_2[of _ _ "(rs ! 0, Stay)"] threeplus2k_3[of _ _ "(rs ! 0, Stay)"] by (cases) simp_all next case b consider "j = 2" | "3 \ j \ j < k + 3" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis using cmdL8_def b threeplus2k_2[of _ _ "(rs ! 0, Stay)"] threeplus2k_3[of _ _ "(rs ! 0, Stay)"] by (cases) simp_all next case c consider "j = 2" | "3 \ j \ j < k + 3" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis using cmdL8_def c threeplus2k_2[of _ _ "(rs ! 0, Stay)"] threeplus2k_3[of _ _ "(rs ! 0, Stay)"] by (cases) simp_all next case d consider "j = 2" | "3 \ j \ j < k + 3" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis using cmdL8_def d threeplus2k_2[of _ _ "(rs ! 0, Stay)"] threeplus2k_3[of _ _ "(rs ! 0, Stay)"] by (cases) simp_all qed qed then have "immobile (tmL78 jj) j (2 * k + 3)" if "jj < k" for jj by (auto simp add: that tmL78_def) then have "immobile (tmL48 jj) j (2 * k + 3)" if "jj < k" for jj using tmL48_def immobile_sequential[OF tmL47_tm tmL78_tm] L47 assms that by simp then have L49: "immobile (tmL49 jj) j (2 * k + 3)" if "jj < k" for jj using tmL49_def immobile_sequential[OF tmL48_tm] tm_left_until1_tm immobile_tm_left_until assms that by simp have L49_upt: "immobile (tmL49_upt j') j (2 * k + 3)" if "j' \ k" for j' using that proof (induction j') case 0 then show ?case by auto next case (Suc j') have "tmL49_upt (Suc j') = tmL49_upt j' ;; tmL49 j'" by simp moreover have "turing_machine (2*k+3) G' (tmL49_upt j')" using tmL49_upt_tm Suc by simp moreover have "immobile (tmL49_upt j') j (2*k+3)" using Suc by simp moreover have "turing_machine (2*k+3) G' (tmL49 j')" using tmL49_tm Suc by simp moreover have "immobile (tmL49 j') j (2*k+3)" using L49 Suc by simp ultimately show ?case using immobile_sequential by simp qed then have "immobile tmL9 j (2 * k + 3)" using tmL9_def immobile_sequential[OF tmL4_tm tmL49_upt_tm] L4 by simp then have L10: "immobile tmL10 j (2 * k + 3)" using tmL10_def immobile_sequential[OF tmL9_tm tmC_tm] C by simp have "cmdL11 rs [~] j = Stay" if "length rs = 2 * k + 3" and "jj < k" for rs jj proof - consider "j = 2" | "3 \ j \ j < 3 + k" | "3 + k \ j \ j < 2 * k + 3" using assms by linarith then show ?thesis proof (cases) case 1 then show ?thesis by (simp add: cmdL11_def) next case 2 then show ?thesis using cmdL11_def threeplus2k_2[where ?a="(rs ! 0, Stay)"] by simp next case 3 then show ?thesis using cmdL11_def threeplus2k_3[where ?a="(rs ! 0, Stay)"] by simp qed qed then have "immobile [cmdL11] j (2 * k + 3)" using k_ge_2 assms by force then have "immobile tmL11 j (2 * k + 3)" using tmL11_def immobile_sequential[OF tmL10_tm tmL1011_tm] L10 by simp then have "immobile tmL12 j (2 * k + 3)" using tmL12_def immobile_sequential[OF tmL11_tm tm_left_until1_tm] immobile_tm_left_until assms by simp then have "immobile tmL13 j (2 * k + 3)" using tmL13_def immobile_sequential[OF tmL12_tm tm_write_many_tm] immobile_tm_write_many assms k_ge_2 G'_ge(1) by simp then have "immobile tmLoop j (2 * k + 3)" using tmLoop_def C immobile_loop[OF tmC_tm tmL13_tm] assms(2) by simp then have "immobile tm10 j (2 * k + 3)" using tm10_def immobile_sequential[OF tm9_tm tmLoop_tm] 9 by simp then have "immobile tm11 j (2 * k + 3)" using tm11_def immobile_sequential[OF tm10_tm tm_ltrans_until_tm] ec1 G'_ge immobile_tm_ltrans_until assms by simp then show "immobile tm12 j (2 * k + 3)" using tm12_def immobile_sequential[OF tm11_tm tm_rtrans_tm] ec1 G'_ge immobile_tm_rtrans assms by simp qed definition "tps12'' zs \ [(\zs\, length zs + 1), (exec zs (Suc (fmt (length zs))) <:> 1, 1)]" lemma tps12'': assumes "bit_symbols zs" shows "tps12'' zs = take 2 (tps12' zs)" using tps12'_def tps12''_def assms by simp text \ This is the actual simulator Turing machine we are constructing in this section. It is @{const tm12} stripped of all memorization tapes: \ definition "tmO2T \ icartesian (2 * k + 1) tm12 G'" lemma tmO2T_tm: "turing_machine 2 G' tmO2T" unfolding tmO2T_def using immobile_tm12 tm12_tm icartesian_tm[of "2 * k + 1" G'] by (metis (no_types, lifting) One_nat_def Suc_le_lessD add.assoc add_less_mono1 le_add2 numeral_3_eq_3 one_add_one plus_1_eq_Suc) text \ The constructed two-tape Turing machine computes the same output as the original Turing machine. \ lemma tmO2T: assumes "bit_symbols zs" shows "traces tmO2T (snd (start_config 2 zs)) (es12 zs) (tps12'' zs)" proof - have *: "2 * k + 1 + 2 = 2 * k + 3" by simp then have "turing_machine (2 * k + 1 + 2) G' tm12" using tm12_tm by metis moreover have "\j. j < 2 * k + 1 \ immobile tm12 (j + 2) (2 * k + 1 + 2)" using immobile_tm12 immobile_def by simp moreover have "\iTime complexity\ text \ This is the bound for the running time of @{const tmO2T}: \ definition TTT :: "nat \ nat" where "TTT \ \n. length (es_fmt n) + 43 * k * Suc (fmt n) * Suc (fmt n)" lemma execute_tmO2T: assumes "bit_symbols zs" shows "execute tmO2T (start_config 2 zs) (TTT (length zs)) = (length tmO2T, tps12'' zs)" proof - have "trace tmO2T (start_config 2 zs) (es12 zs) (length tmO2T, tps12'' zs)" using tmO2T assms traces_def start_config_def by simp then have "execute tmO2T (start_config 2 zs) (length (es12 zs)) = (length tmO2T, tps12'' zs)" using trace_def by simp moreover have "length (es12 zs) \ TTT (length zs)" using assms length_es12 TTT_def by simp ultimately show ?thesis by (metis execute_after_halting_ge fst_conv) qed text \ The simulator TM @{const tmO2T} halts with the output tape head on cell~1. \ lemma execute_tmO2T_1: assumes "bit_symbols zs" shows "execute tmO2T (start_config 2 zs) (TTT (length zs)) 1 = (execute M (start_config k zs) (T (length zs)) <:> 1, 1)" proof - have "T (length zs) \ Suc (fmt (length zs))" by (metis add_leD1 le_Suc_eq fmt(4) T'_def) then have *: "execute M (start_config k zs) (T (length zs)) = execute M (start_config k zs) (Suc (fmt (length zs)))" using execute_after_halting_ge time_bound_T time_bound_def assms by (metis (no_types, lifting)) have "execute tmO2T (start_config 2 zs) (TTT (length zs)) = (length tmO2T, tps12'' zs)" using assms execute_tmO2T by simp then have "execute tmO2T (start_config 2 zs) (TTT (length zs)) 1 = (execute M (start_config k zs) (Suc (fmt (length zs))) <:> 1, 1)" using tps12''_def exec_def assms by simp then show ?thesis using * by simp qed lemma poly_TTT: "big_oh_poly TTT" proof - have 1: "big_oh_poly (\n. length (es_fmt n))" using fmt(2) by simp have "big_oh_poly (\n. fmt n + 1)" using fmt(3) big_oh_poly_const big_oh_poly_sum by blast then have "big_oh_poly (\n. Suc (fmt n))" by simp then have "big_oh_poly (\n. Suc (fmt n) * Suc (fmt n))" using big_oh_poly_prod by blast moreover have "big_oh_poly (\n. 43 * k)" using big_oh_poly_const by simp ultimately have "big_oh_poly (\n. 43 * k * (Suc (fmt n) * Suc (fmt n)))" using big_oh_poly_prod by blast moreover have "(\n. 43 * k * (Suc (fmt n) * Suc (fmt n))) = (\n. 43 * k * Suc (fmt n) * Suc (fmt n))" by (metis (mono_tags, opaque_lifting) mult.assoc) ultimately have "big_oh_poly (\n. 43 * k * Suc (fmt n) * Suc (fmt n))" by simp then have "big_oh_poly (\n. length (es_fmt n) + 43 * k * Suc (fmt n) * Suc (fmt n))" using 1 big_oh_poly_sum by simp then show ?thesis unfolding TTT_def by simp qed subsection \Obliviousness\ text \ The two-tape simulator machine is oblivious. \ lemma tmO2T_oblivious: assumes "length zs1 = length zs2" and "bit_symbols zs1" and "bit_symbols zs2" shows "es12 zs1 = es12 zs2" proof - have "es1 zs1 = es1 zs2" using es1_def assms by simp moreover have "es1_2 zs1 = es1_2 zs2" using es1_2_def assms by (metis (no_types, lifting)) ultimately have "es2 zs1 = es2 zs2" using es2_def assms by simp then have "es3 zs1 = es3 zs2" using es3_def assms by simp then have "es4 zs1 = es4 zs2" using es4_def assms by simp then have "es5 zs1 = es5 zs2" using es5_def assms by simp then have "es6 zs1 = es6 zs2" using es6_def assms by simp then have "es7 zs1 = es7 zs2" using es7_def assms by simp then have "es8 zs1 = es8 zs2" using es8_def assms by simp then have 9: "es9 zs1 = es9 zs2" using es9_def assms by simp have C: "esC zs1 t = esC zs2 t" for t using esC_def assms by simp then have Loop_break: "esLoop_break zs1 = esLoop_break zs2" using esLoop_break_def tpsC1_def tpsL_def assms by simp have "esL1 zs1 = esL1 zs2" using esL1_def assms by auto moreover have "esL1_2 zs1 = esL1_2 zs2" using esL1_2_def assms by simp ultimately have "esL2 zs1 = esL2 zs2" using esL2_def assms by auto then have "esL3 zs1 = esL3 zs2" using esL3_def assms by auto then have L4: "esL4 zs1 = esL4 zs2" using esL4_def assms by auto have "esL45 zs1 = esL45 zs2" using esL45_def assms by simp then have "esL46 zs1 = esL46 zs2" using esL46_def assms by simp moreover have "esL67 zs1 = esL67 zs2" using esL67_def assms by simp ultimately have "esL47 zs1 = esL47 zs2" using esL47_def assms by simp moreover have "esL78 zs1 = esL78 zs2" using esL78_def assms by simp ultimately have "esL48 zs1 = esL48 zs2" using esL48_def assms by simp then have "esL49 zs1 = esL49 zs2" using esL49_def assms by simp then have "esL49_upt zs1 = esL49_upt zs2" using esL49_upt_def assms by simp then have "esL9 zs1 = esL9 zs2" using esL9_def L4 assms by auto then have "esL10 zs1 = esL10 zs2" using esL10_def C assms by auto then have "esL11 zs1 = esL11 zs2" using esL11_def assms by auto then have "esL12 zs1 = esL12 zs2" using esL12_def assms by auto then have "esL13 zs1 = esL13 zs2" using esL13_def assms by auto then have "esLoop_while zs1 = esLoop_while zs2" using esLoop_while_def C tpsC1_def tpsL13_def tpsL_def assms by auto then have "esLoop zs1 = esLoop zs2" using esLoop_def Loop_break assms by auto then have "es10 zs1 = es10 zs2" using es10_def 9 assms by auto then have "es11 zs1 = es11 zs2" using es11_def assms by simp then show "es12 zs1 = es12 zs2" using es12_def assms by simp qed end (* locale two_tape *) section \$\NP$ and obliviousness\label{s:oblivious-np}\ text \ This section presents the main result of this chapter: For every language $L \in \NP$ there is a polynomial-time two-tape oblivious verifier TM that halts with the output tape head on a \textbf{1} symbol iff.\ in the input $\langle x, u\rangle$, the string $u$ is a certificate for $x$. The proof combines two lemmas. First @{thm [source] NP_output_len_1}, which says that we can assume the verifier outputs only one symbol (namely, \textbf{0} or \textbf{1}), and second @{thm [source] two_tape.execute_tmO2T_1}, which says that the two-tape oblivious TM halts with output tape head in cell~1. This cell will contain either \textbf{0} or \textbf{1} by the first lemma. \ lemma NP_imp_oblivious_2tape: fixes L :: language assumes "L \ \\

" obtains M G T p where "big_oh_poly T" and "polynomial p" and "turing_machine 2 G M" and "oblivious M" and "\y. bit_symbols y \ fst (execute M (start_config 2 y) (T (length y))) = length M" and "\x. x \ L \ (\u. length u = p (length x) \ execute M (start_config 2 \x; u\) (T (length \x; u\)) <.> 1 = \)" proof - define Q where "Q = (\L k G M p T fverify. turing_machine k G M \ polynomial p \ big_oh_poly T \ computes_in_time k M fverify T \ (\y. length (fverify y) = 1) \ (\x. (x \ L) = (\u. length u = p (length x) \ fverify \x, u\ = [\])))" have "\\

= {L :: language. \k G M p T fverify. Q L k G M p T fverify}" unfolding NP_output_len_1 Q_def by simp then obtain k G M p T fverify where "Q L k G M p T fverify" using assms by blast then have alt: "turing_machine k G M" "polynomial p" "big_oh_poly T" "computes_in_time k M fverify T" "\y. length (fverify y) = 1" "\x. (x \ L) = (\u. length u = p (length x) \ fverify \x, u\ = [\])" using Q_def by simp_all have tm_M: "turing_machine k G M" using alt(1) . have poly_T: "big_oh_poly T" using alt(3) . have time_bound_T: "time_bound M k T" unfolding time_bound_def proof standard+ fix zs assume zs: "bit_symbols zs" define x where "x = symbols_to_string zs" then have "zs = string_to_symbols x" using bit_symbols_to_symbols zs by simp then show "fst (execute M (start_config k zs) (T (length zs))) = length M" using computes_in_time_def alt(4) by (metis (no_types, lifting) execute_after_halting_ge length_map running_timeD(1)) qed interpret two: two_tape M k G T using tm_M poly_T time_bound_T two_tape_def by simp let ?M = two.tmO2T let ?T = two.TTT let ?G = two.G' have "big_oh_poly ?T" using two.poly_TTT . moreover have "polynomial p" using alt(2) . moreover have "turing_machine 2 ?G ?M" using two.tmO2T_tm . moreover have "oblivious ?M" proof - let ?e = "\n. two.es12 (replicate n 2)" have "\tps. trace ?M (start_config 2 zs) (?e (length zs)) (length ?M, tps)" if "bit_symbols zs" for zs proof - have "traces ?M (snd (start_config 2 zs)) (two.es12 zs) (two.tps12'' zs)" using that two.tmO2T by simp then have *: "trace ?M (start_config 2 zs) (two.es12 zs) (length ?M, two.tps12'' zs)" using start_config_def traces_def by simp let ?r = "replicate (length zs) 2" have "length zs = length ?r" by simp then have "two.es12 zs = two.es12 ?r" using two.tmO2T_oblivious that by (metis nth_replicate) then have "trace ?M (start_config 2 zs) (?e (length zs)) (length ?M, two.tps12'' zs)" using * by simp then show ?thesis by auto qed then show ?thesis using oblivious_def by fast qed moreover have "fst (execute ?M (start_config 2 y) (?T (length y))) = length ?M" if "bit_symbols y" for y using that two.execute_tmO2T by simp moreover have "x \ L \ (\u. length u = p (length x) \ execute ?M (start_config 2 \x; u\) (?T (length \x; u\)) <.> 1 = \)" (is "?lhs = ?rhs") for x proof show "?lhs \ ?rhs" proof - assume ?lhs then have "\u. length u = p (length x) \ fverify \x, u\ = [\]" using alt(6) by simp then obtain u where u: "length u = p (length x)" "fverify \x, u\ = [\]" by auto let ?y = "fverify \x, u\" let ?cfg = "execute M (start_config k \x; u\) (T (length \x, u\))" have "computes_in_time k M fverify T" using alt(4) by simp then have cfg: "?cfg <:> 1 = string_to_contents ?y" using computes_in_time_execute by simp have "bit_symbols \x; u\" by simp then have "execute ?M (start_config 2 \x; u\) (?T (length \x; u\)) 1 = (execute M (start_config k \x; u\) (T (length \x; u\)) <:> 1, 1)" using two.execute_tmO2T_1 by blast then have "execute ?M (start_config 2 \x; u\) (?T (length \x; u\)) 1 = (string_to_contents ?y, 1)" using cfg by simp then have "execute ?M (start_config 2 \x; u\) (?T (length \x; u\)) 1 = (string_to_contents [\], 1)" using u(2) by auto moreover have "|.| (string_to_contents [\], 1) = \" by simp ultimately have "execute ?M (start_config 2 \x; u\) (?T (length \x; u\)) <.> 1 = \" by simp then show ?thesis using u(1) by auto qed show "?rhs \ ?lhs" proof - assume ?rhs then obtain u where u: "length u = p (length x)" "execute ?M (start_config 2 \x; u\) (?T (length \x; u\)) <.> 1 = \" by auto let ?zs = "\x; u\" have "bit_symbols \x; u\" by simp then have *: "execute ?M (start_config 2 ?zs) (?T (length ?zs)) 1 = (execute M (start_config k ?zs) (T (length ?zs)) <:> 1, 1)" using two.execute_tmO2T_1 by blast let ?y = "fverify \x, u\" let ?cfg = "execute M (start_config k ?zs) (T (length \x, u\))" have "computes_in_time k M fverify T" using alt(4) by simp then have cfg: "?cfg <:> 1 = string_to_contents ?y" using computes_in_time_execute by simp then have "execute ?M (start_config 2 ?zs) (?T (length ?zs)) 1 = (string_to_contents (fverify \x, u\), 1)" using * by simp then have "execute ?M (start_config 2 ?zs) (?T (length ?zs)) <.> 1 = string_to_contents (fverify \x, u\) 1" by simp then have **: "string_to_contents (fverify \x, u\) 1 = \" using u(2) by simp have "length (fverify \x, u\) = 1" using alt(5) by simp then have "string_to_contents (fverify \x, u\) 1 = (if fverify \x, u\ ! 0 then \ else \)" by simp then have "(if fverify \x, u\ ! 0 then \ else \) = \" using ** by auto then have "fverify \x, u\ ! 0 = \" by (metis numeral_eq_iff semiring_norm(89)) moreover have "y = [\]" if "length y = 1" and "y ! 0" for y using that by (metis (full_types) One_nat_def Suc_length_conv length_0_conv nth_Cons') ultimately have "fverify \x, u\ = [\]" using alt(5) by simp then show ?lhs using u(1) alt(6) by auto qed qed ultimately show ?thesis using that by simp qed end \ No newline at end of file diff --git a/thys/Cook_Levin/Oblivious_Polynomial.thy b/thys/Cook_Levin/Oblivious_Polynomial.thy --- a/thys/Cook_Levin/Oblivious_Polynomial.thy +++ b/thys/Cook_Levin/Oblivious_Polynomial.thy @@ -1,1473 +1,1473 @@ section \Constructing polynomials in polynomial time\label{s:oblivious-polynomial}\ theory Oblivious_Polynomial imports Oblivious begin text \ Our current goal is to simulate a polynomial time multi-tape TM on a two-tape oblivious TM in polynomial time. To help with the obliviousness we first want to mark on the simulator's output tape a space that is at least as large as the space the simulated TM uses on any of its tapes but that still is only polynomial in size. In this section we construct oblivious Turing machines for that. An upper bound for the size this space is provided by the simulated TM's running time, which by assumption is polynomially bounded. So for our purposes any polynomially bounded function that bounds the running time will do. In this section we devise a family of two-tape oblivious TMs that contains for every polynomial degree $d \ge 1$ a TM that writes $\mathbf{1}^{p(n)}$ to the output tape for some polynomial $p$ with $p(n) \ge n^d$, where $n$ is the length of the input to the TM. Together with a TM that appends a constant number $c$ of ones we get a family of TMs, parameterized by $c$ and $d$, that runs in polynomial time and writes more than $c + n^d$ symbols~\textbf{1} to the work tape. This meets our goal for this section because every polynomially bounded function is bounded by a function of the form $n \mapsto c + n^d$ for some $c, d\in\nat$. The TMs in the family are built from three building block TMs. The first TM initializes its output tape with $\mathbf{1}^n$ where $n$ is the length of the input. The second TM multiplies the number of symbols on the output tape by (roughly) the length of the input, turning a sequence $\mathbf{1}^m$ into (roughly) $\mathbf{1}^{mn}$ for arbitrary $m$. The third TM appends $\mathbf{1}^c$ for some constant $c$. By repeating the second TM we can achieve arbitrarily high polynomial degrees. All three TMs do essentially only one thing with the input, namely copying it to the output tape while mapping all symbols to \textbf{1}, which is an operation that depends only on the length of the input. Therefore all three TMs are oblivious, and combining them also yields an oblivious TM. The Turing machines require one extra symbol beyond the four default symbols, but work for all alphabet sizes $G \ge 5$. \null \ locale turing_machine_poly = fixes G :: nat assumes G: "G \ 5" begin lemma G_ge4 [simp]: "G \ 4" using G by linarith abbreviation tps_ones :: "symbol list \ nat \ tape list" where "tps_ones zs m \ [(\zs\, 1), (\i. if i = 0 then \ else if i \ m then \ else \, 1)]" subsection \Initializing the output tape\ text \ The first building block is a TM that ``copies'' the input to the output tape, thereby replacing every symbol by the symbol \textbf{1}. \ definition tmA :: machine where "tmA \ tm_right 0 ;; tm_right 1 ;; tm_const_until 0 1 {\} \ ;; tm_cr 0 ;; tm_cr 1" lemma tmA_tm: "turing_machine 2 G tmA" unfolding tmA_def using tm_right_tm tm_const_until_tm tm_cr_tm G by simp definition "tm1 \ tm_right 0" definition "tm2 \ tm1 ;; tm_right 1" definition "tm3 \ tm2 ;; tm_const_until 0 1 {\} \" definition "tm4 \ tm3 ;; tm_cr 0" definition "tm5 \ tm4 ;; tm_cr 1" lemma tm5_eq_tmA: "tm5 = tmA" unfolding tmA_def tm5_def tm4_def tm3_def tm2_def tm1_def by simp definition tps0 :: "symbol list \ tape list" where "tps0 zs \ [(\zs\, 0), (\i. if i = 0 then \ else \, 0)]" lemma length_tps0 [simp]: "length (tps0 n) = 2" unfolding tps0_def by simp definition tps1 :: "symbol list \ tape list" where "tps1 zs \ [(\zs\, 1), (\i. if i = 0 then \ else \, 0)]" definition es1 :: "(nat \ nat) list" where "es1 \ [(1, 0)]" lemma tm1: "traces tm1 (tps0 zs) es1 (tps1 zs)" unfolding tm1_def by (rule traces_tm_right_0I) (simp_all add: tps0_def tps1_def es1_def) definition tps2 :: "symbol list \ tape list" where "tps2 zs \ [(\zs\, 1), (\i. if i = 0 then \ else \, 1)]" definition es2 :: "(nat \ nat) list" where "es2 \ es1 @ [(1, 1)]" lemma length_es2: "length es2 = 2" unfolding es1_def es2_def by simp lemma tm2: "traces tm2 (tps0 zs) es2 (tps2 zs)" unfolding tm2_def es2_def proof (rule traces_sequential[OF tm1]) show "traces (tm_right 1) (tps1 zs) [(1, 1)] (tps2 zs)" using tps1_def tps2_def by (intro traces_tm_right_1I) simp_all qed definition tps3 :: "symbol list \ tape list" where "tps3 zs \ [(\zs\, length zs + 1), (\i. if i = 0 then \ else if i \ length zs then \ else \, length zs + 1)]" definition es23 :: "nat \ (nat \ nat) list" where "es23 n \ map (\i. (i + 2, i + 2)) [0.. (nat \ nat) list" where "es3 n \ es2 @ (es23 n)" lemma length_es3: "length (es3 n) = n + 3" unfolding es3_def es23_def using length_es2 by simp lemma tm3: assumes "proper_symbols zs" shows "traces tm3 (tps0 zs) (es3 (length zs)) (tps3 zs)" unfolding tm3_def es3_def proof (rule traces_sequential[OF tm2]) show "traces (tm_const_until 0 1 {\} \) (tps2 zs) (es23 (length zs)) (tps3 zs)" proof (rule traces_tm_const_until_01I) show "1 < length (tps2 zs)" using tps2_def by simp show "rneigh (tps2 zs ! 0) {\} (length zs)" using rneigh_def contents_def tps2_def assms by auto show "es23 (length zs) = map (\i. (tps2 zs :#: 0 + Suc i, tps2 zs :#: 1 + Suc i)) [0.. (length zs)]" using tps3_def tps2_def constplant by auto qed qed definition tps4 :: "symbol list \ tape list" where "tps4 zs \ [(\zs\, 1), (\i. if i = 0 then \ else if i \ length zs then \ else \, length zs + 1)]" definition es34 :: "nat \ (nat \ nat) list" where "es34 n \ map (\i. (i, n + 1)) (rev [0.. (nat \ nat) list" where "es4 n \ es3 n @ es34 n" lemma length_es4: "length (es4 n) = 2 * n + 6" unfolding es4_def es34_def using length_es3 by simp lemma tm4: assumes "proper_symbols zs" shows "traces tm4 (tps0 zs) (es4 (length zs)) (tps4 zs)" unfolding tm4_def es4_def proof (rule traces_sequential[OF tm3]) show "proper_symbols zs" using assms . show "traces (tm_cr 0) (tps3 zs) (es34 (length zs)) (tps4 zs)" proof (rule traces_tm_cr_0I) show "1 < length (tps3 zs)" using tps3_def by simp show "clean_tape (tps3 zs ! 0)" using assms tps3_def by simp show "es34 (length zs) = map (\i. (i, tps3 zs :#: 1)) (rev [0.. nat \ tape list" where "tps5 zs m \ tps_ones zs m" definition es45 :: "nat \ (nat \ nat) list" where "es45 n \ map (\i. (1, i)) (rev [0.. (nat \ nat) list" where "es5 n \ es4 n @ es45 n" lemma length_es5: "length (es5 n) = 3 * n + 9" unfolding es5_def es45_def using length_es4 by simp lemma tm5: assumes "proper_symbols zs" shows "traces tm5 (tps0 zs) (es5 (length zs)) (tps5 zs (length zs))" unfolding tm5_def es5_def proof (rule traces_sequential[OF tm4]) show "proper_symbols zs" using assms . show "traces (tm_cr 1) (tps4 zs) (es45 (length zs)) (tps5 zs (length zs))" proof (rule traces_tm_cr_1I) show "1 < length (tps4 zs)" using tps4_def by simp show "clean_tape (tps4 zs ! 1)" using tps4_def clean_tape_def by simp show "es45 (length zs) = map (Pair (tps4 zs :#: 0)) (rev [0..Multiplying by the input length\ text \ The next Turing machine turns a symbol sequence $\mathbf{1}^m$ on its output tape into $\mathbf{1}^{m+1+mn}$ where $n$ is the length of the input. The TM first appends to the output tape symbols a @{text "\"} symbol. Then it performs a loop with $m$ iterations. In each iteration it replaces a @{text \} before the @{text "\"} by @{text \} in order to count the iterations. Then it copies (replacing each symbol by @{text \}) the input after the output tape symbols. After the loop the output tape contains $\mathbf{0}^m|\mathbf{1}^{mn}$. Finally the @{text "\"} and @{text \} symbols are replaced by @{text \} symbols, yielding the desired result. \ definition tmB :: machine where "tmB \ tm_right_until 1 {\} ;; tm_write 1 \ ;; tm_cr 1 ;; WHILE tm_right_until 1 {\, \} ; \rs. rs ! 1 = \ DO tm_write 1 \ ;; tm_right_until 1 {0} ;; tm_const_until 0 1 {\} \ ;; tm_cr 0 ;; tm_cr 1 DONE ;; tm_write 1 \ ;; tm_cr 1 ;; tm_const_until 1 1 {\} \ ;; tm_cr 1" lemma tmB_tm: "turing_machine 2 G tmB" unfolding tmB_def using tm_right_until_tm tm_write_tm tm_cr_tm tm_const_until_tm G turing_machine_loop_turing_machine turing_machine_sequential_turing_machine by simp definition "tmB1 \ tm_right_until 1 {\}" definition "tmB2 \ tmB1 ;; tm_write 1 \" definition "tmB3 \ tmB2 ;; tm_cr 1" definition "tmK1 \ tm_right_until 1 {\, \}" definition "tmL1 \ tm_write 1 \" definition "tmL2 \ tmL1 ;; tm_right_until 1 {\}" definition "tmL3 \ tmL2 ;; tm_const_until 0 1 {\} \" definition "tmL4 \ tmL3 ;; tm_cr 0" definition "tmL5 \ tmL4 ;; tm_cr 1" definition "tmLoop \ WHILE tmK1 ; \rs. rs ! 1 = \ DO tmL5 DONE" definition "tmB4 \ tmB3 ;; tmLoop" definition "tmB5 \ tmB4 ;; tm_write 1 \" definition "tmB6 \ tmB5 ;; tm_cr 1" definition "tmB7 \ tmB6 ;; tm_const_until 1 1 {\} \" definition "tmB8 \ tmB7 ;; tm_cr 1" lemma tmB_eq_tmB8: "tmB = tmB8" unfolding tmB_def tmB1_def tmB2_def tmB3_def tmK1_def tmL1_def tmL2_def tmL3_def tmL4_def tmL5_def tmLoop_def tmB4_def tmB5_def tmB6_def tmB7_def tmB8_def by simp definition tpsB1 :: "symbol list \ nat \ tape list" where "tpsB1 zs m \ [(\zs\, 1), (\i. if i = 0 then \ else if i \ m then \ else \, m + 1)]" definition esB1 :: "nat \ nat \ (nat \ nat) list" where "esB1 n m \ map (\i. (1, 1 + Suc i)) [0..i. (tps_ones zs m :#: 0, tps_ones zs m :#: 1 + Suc i)) [0.. nat \ tape list" where "tpsB2 zs m \ [(\zs\, 1), (\i. if i = 0 then \ else if i \ m then \ else if i = m + 1 then \ else \, m + 1)]" definition esB12 :: "nat \ nat \ (nat \ nat) list" where "esB12 n m \ [(1, m + 1)]" definition esB2 :: "nat \ nat \ (nat \ nat) list" where "esB2 n m \ esB1 n m @ esB12 n m" lemma length_esB2: "length (esB2 n m) = m + 2" by (simp add: esB12_def esB2_def length_esB1) lemma tmB2: assumes "proper_symbols zs" shows "traces tmB2 (tps_ones zs m) (esB2 (length zs) m) (tpsB2 zs m)" unfolding tmB2_def esB2_def proof (rule traces_sequential[OF tmB1]) show "proper_symbols zs" using assms . show "traces (tm_write 1 \) (tpsB1 zs m) (esB12 (length zs) m) (tpsB2 zs m)" proof (rule traces_tm_write_1I) show "1 < length (tpsB1 zs m)" using tpsB1_def by simp_all show "esB12 (length zs) m = [(tpsB1 zs m :#: 0, tpsB1 zs m :#: 1)]" using tpsB1_def by (simp add: esB12_def) show "tpsB2 zs m = (tpsB1 zs m)[1 := tpsB1 zs m ! 1 |:=| \]" using tpsB2_def tpsB1_def by auto qed qed definition tpsB3 :: "symbol list \ nat \ tape list" where "tpsB3 zs m \ [(\zs\, 1), (\i. if i = 0 then \ else if i \ m then \ else if i = m + 1 then \ else 0, 1)]" definition esB23 :: "nat \ nat \ (nat \ nat) list" where "esB23 n m \ map (Pair 1) (rev [0.. nat \ (nat \ nat) list" where "esB3 n m \ esB2 n m @ esB23 n m" lemma length_esB3: "length (esB3 n m) = 2 * m + 5" by (simp add: esB3_def length_esB2 esB23_def) lemma tmB3: assumes "proper_symbols zs" shows "traces tmB3 (tps_ones zs m) (esB3 (length zs) m) (tpsB3 zs m)" unfolding tmB3_def esB3_def proof (rule traces_sequential[OF tmB2]) show "proper_symbols zs" using assms . show "traces (tm_cr 1) (tpsB2 zs m) (esB23 (length zs) m) (tpsB3 zs m)" proof (rule traces_tm_cr_1I) show "1 < length (tpsB2 zs m)" using tpsB2_def by simp show "clean_tape (tpsB2 zs m ! 1)" using tpsB2_def clean_tape_def by simp show "esB23 (length zs) m = map (Pair (tpsB2 zs m :#: 0)) (rev [0.. nat \ nat \ tape list" where "tpsK0 zs m i \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ i then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + i * length zs then \ else 0, 1)]" lemma tpsK0_eq_tpsB3: "tpsK0 zs m 0 = tpsB3 zs m" using tpsK0_def tpsB3_def by auto definition tpsK1 :: "symbol list \ nat \ nat \ tape list" where "tpsK1 zs m i \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ i then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + i * length zs then \ else 0, i + 1)]" definition esK1 :: "nat \ nat \ nat \ (nat \ nat) list" where "esK1 n m i \ map (\i. (1, 1 + Suc i)) [0.., \} i" using tpsK0_def rneigh_def assms(2) by simp show "esK1 (length zs) m i = map (\j. (tpsK0 zs m i :#: 0, tpsK0 zs m i :#: 1 + Suc j)) [0.. nat \ nat \ tape list" where "tpsL1 zs m i \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ i + 1 then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + i * length zs then \ else 0, i + 1)]" definition esL1 :: "nat \ nat \ nat \ (nat \ nat) list" where "esL1 n m i \ [(1, i + 1)]" lemma tmL1: assumes "proper_symbols zs" shows "traces tmL1 (tpsK1 zs m i) (esL1 (length zs) m i) (tpsL1 zs m i)" unfolding tmL1_def using G by (intro traces_tm_write_1I) (auto simp add: tpsL1_def tpsK1_def esL1_def) definition tpsL2 :: "symbol list \ nat \ nat \ tape list" where "tpsL2 zs m i \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ i + 1 then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + i * length zs then \ else 0, m + 2 + i * length zs)]" definition esL12 :: "nat \ nat \ nat \ (nat \ nat) list" where "esL12 n m i \ map (\j. (1, i + 1 + Suc j)) [0.. nat \ nat \ (nat \ nat) list" where "esL2 n m i \ esL1 n m i @ esL12 n m i" lemma length_esL2: "i < m \ length (esL2 n m i) = 3 + m - i + i * n" by (auto simp add: esL2_def esL12_def esL1_def) text \ A simplified upper bound for the running time: \ corollary length_esL2': "i < m \ length (esL2 n m i) \ 3 + m + i * n" by (simp add: length_esL2) lemma tmL2: assumes "proper_symbols zs" and "i < m" shows "traces tmL2 (tpsK1 zs m i) (esL2 (length zs) m i) (tpsL2 zs m i)" unfolding tmL2_def esL2_def proof (rule traces_sequential[OF tmL1]) show "proper_symbols zs" using assms(1) . show "traces (tm_right_until 1 {0}) (tpsL1 zs m i) (esL12 (length zs) m i) (tpsL2 zs m i)" proof (rule traces_tm_right_until_1I) show "1 < length (tpsL1 zs m i)" using tpsL1_def by simp show "rneigh (tpsL1 zs m i ! 1) {0} (m + 2 + i * length zs - (i + 1))" using rneigh_def assms(2) by (auto simp add: tpsL1_def) show "esL12 (length zs) m i = map (\j. (tpsL1 zs m i :#: 0, tpsL1 zs m i :#: 1 + Suc j)) [0.. nat \ nat \ tape list" where "tpsL3 zs m i \ [(\zs\, length zs + 1), (\x. if x = 0 then \ else if x \ i + 1 then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + (i + 1) * length zs then \ else 0, m + 2 + (i + 1) * length zs)]" definition esL23 :: "nat \ nat \ nat \ (nat \ nat) list" where "esL23 n m i \ map (\j. (1 + Suc j, m + 2 + i * n + Suc j)) [0.. nat \ nat \ (nat \ nat) list" where "esL3 n m i \ esL2 n m i @ esL23 n m i" lemma length_esL3: "i < m \ length (esL3 n m i) \ 4 + m + (i + 1) * n" by (auto simp add: esL3_def esL23_def) (metis group_cancel.add1 length_esL2') lemma tmL3: assumes "proper_symbols zs" and "i < m" shows "traces tmL3 (tpsK1 zs m i) (esL3 (length zs) m i) (tpsL3 zs m i)" unfolding tmL3_def esL3_def proof (rule traces_sequential[OF tmL2]) show "proper_symbols zs" and "i < m" using assms . show "traces (tm_const_until 0 1 {\} \) (tpsL2 zs m i) (esL23 (length zs) m i) (tpsL3 zs m i)" proof (rule traces_tm_const_until_01I) show "1 < length (tpsL2 zs m i)" using tpsL2_def by simp show "rneigh (tpsL2 zs m i ! 0) {0} (length zs)" using assms(1) rneigh_def contents_def by (auto simp add: tpsL2_def) show "esL23 (length zs) m i = map (\j. (tpsL2 zs m i :#: 0 + Suc j, tpsL2 zs m i :#: 1 + Suc j)) [0.. (length zs)]" using constplant assms by (auto simp add: tpsL2_def tpsL3_def) qed qed definition tpsL4 :: "symbol list \ nat \ nat \ tape list" where "tpsL4 zs m i \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ i + 1 then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + (i + 1) * length zs then \ else 0, m + 2 + (i + 1) * length zs)]" definition esL34 :: "nat \ nat \ nat \ (nat \ nat) list" where "esL34 n m i \ map (\j. (j, m + 2 + (i + 1) * n)) (rev [0.. length (esL34 n m i) = n + 3" unfolding esL34_def by simp definition esL4 :: "nat \ nat \ nat \ (nat \ nat) list" where "esL4 n m i \ esL3 n m i @ esL34 n m i" lemma length_esL4: "i < m \ length (esL4 n m i) \ 7 + m + (i + 2) * n" using length_esL3 length_esL34 by (auto simp add: esL4_def) (metis ab_semigroup_add_class.add_ac(1) group_cancel.add2) lemma tmL4: assumes "proper_symbols zs" and "i < m" shows "traces tmL4 (tpsK1 zs m i) (esL4 (length zs) m i) (tpsL4 zs m i)" unfolding tmL4_def esL4_def proof (rule traces_sequential[OF tmL3]) show "proper_symbols zs" "i < m" using assms . show "traces (tm_cr 0) (tpsL3 zs m i) (esL34 (length zs) m i) (tpsL4 zs m i)" proof (rule traces_tm_cr_0I) show "1 < length (tpsL3 zs m i)" using tpsL3_def by simp show "clean_tape (tpsL3 zs m i ! 0)" using tpsL3_def assms(1) by simp show "esL34 (length zs) m i = map (\j. (j, tpsL3 zs m i :#: 1)) (rev [0.. nat \ nat \ tape list" where "tpsL5 zs m i \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ i + 1 then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + (i + 1) * length zs then \ else 0, 1)]" definition esL45 :: "nat \ nat \ nat \ (nat \ nat) list" where "esL45 n m i \ map (Pair 1) (rev [0.. nat \ nat \ (nat \ nat) list" where "esL5 n m i \ esL4 n m i @ esL45 n m i" lemma length_esL5: "i < m \ length (esL5 n m i) \ 11 + 2 * m + (2 * i + 3) * n" proof - assume a: "i < m" have "length (esL5 n m i) = length (esL4 n m i) + length (esL45 n m i)" using esL5_def by simp also have "... \ 7 + m + (i + 2) * n + length (esL45 n m i)" using length_esL4[OF a] by simp also have "... = 7 + m + (i + 2) * n + (m + 2 + (i + 1) * n + 2)" using esL45_def by simp also have "... = 11 + 2 * m + (2 * i + 3) * n" by algebra finally show ?thesis . qed lemma tmL5: assumes "proper_symbols zs" and "i < m" shows "traces tmL5 (tpsK1 zs m i) (esL5 (length zs) m i) (tpsL5 zs m i)" unfolding tmL5_def esL5_def proof (rule traces_sequential[OF tmL4]) show "proper_symbols zs" "i < m" using assms . show "traces (tm_cr 1) (tpsL4 zs m i) (esL45 (length zs) m i) (tpsL5 zs m i)" proof (rule traces_tm_cr_1I) show "1 < length (tpsL4 zs m i)" using tpsL4_def by simp show "clean_tape (tpsL4 zs m i ! 1)" using tpsL4_def assms clean_tapeI by simp show "esL45 (length zs) m i = map (Pair (tpsL4 zs m i :#: 0)) (rev [0.. nat \ nat \ (nat \ nat) list" where "esLoop_do n m i \ esK1 n m i @ [(1, i + 1)] @ esL5 n m i @ [(1, 1)]" text \ Using $i < m$ we get this upper bound for the running time of an iteration independent of $i$. \ lemma length_esLoop_do: "i < m \ length (esLoop_do n m i) \ 14 + 3 * m + (2 * m + 3) * n" proof - assume "i < m" have "length (esLoop_do n m i) = length (esK1 n m i) + length (esL5 n m i) + 2" unfolding esLoop_do_def by simp also have "... = i + 3 + (length (esL5 n m i))" using length_esK1 by simp also have "... \ i + 14 + 2 * m + (2 * i + 3) * n" using length_esL5[OF `i < m`] by (simp add: add.assoc) also have "... \ 14 + 3 * m + (2 * i + 3) * n" using `i < m` by simp also have "... \ 14 + 3 * m + (2 * m + 3) * n" using `i < m` by simp finally show ?thesis . qed lemma tmLoop_do: assumes "proper_symbols zs" and "i < m" shows "trace tmLoop (0, tpsK0 zs m i) (esLoop_do (length zs) m i) (0, tpsL5 zs m i)" unfolding tmLoop_def proof (rule tm_loop_sem_true_tracesI[OF tmK1 tmL5]) show "proper_symbols zs" "proper_symbols zs" "i < m" "i < m" using assms by simp_all show "read (tpsK1 zs m i) ! 1 = \" using tpsK1_def assms(2) read_def by simp show "esLoop_do (length zs) m i = esK1 (length zs) m i @ [(tpsK1 zs m i :#: 0, tpsK1 zs m i :#: 1)] @ esL5 (length zs) m i @ [(tpsL5 zs m i :#: 0, tpsL5 zs m i :#: 1)]" by (simp add: esLoop_do_def esK1_def tpsK1_def esL5_def tpsL5_def) qed lemma tpsL5_eq_tpsK0: assumes "proper_symbols zs" and "i < m" shows "tpsL5 zs m i = tpsK0 zs m (Suc i)" using assms tpsL5_def tpsK0_def by auto lemma tmLoop_iteration: assumes "proper_symbols zs" and "i < m" shows "trace tmLoop (0, tpsK0 zs m i) (esLoop_do (length zs) m i) (0, tpsK0 zs m (Suc i))" using assms tmLoop_do tpsL5_eq_tpsK0 by simp definition esLoop_done :: "nat \ nat \ (nat \ nat) list" where "esLoop_done n m \ concat (map (esLoop_do n m) [0.. m * (14 + 3 * m + (2 * m + 3) * n)" using length_concat_le[OF length_esLoop_do] esLoop_done_def by simp definition tpsK_break :: "symbol list \ nat \ tape list" where "tpsK_break zs m \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m then \ else if x = m + 1 then \ else if x \ m + 1 + m * length zs then \ else 0, m + 1)]" definition esK_break :: "nat \ nat \ (nat \ nat) list" where "esK_break n m \ map (\i. (1, 1 + Suc i)) [0.., \} m" using rneigh_def by (simp add: tpsK0_def) show "esK_break (length zs) m = map (\j. (tpsK0 zs m m :#: 0, tpsK0 zs m m :#: 1 + Suc j)) [0.. nat \ (nat \ nat) list" where "esLoop_break n m \ esK_break n m @ [(1, m + 1)]" lemma length_esLoop_break: "length (esLoop_break n m) = m + 2" unfolding esLoop_break_def using length_esK_break by simp lemma tmLoop_break: assumes "proper_symbols zs" shows "traces tmLoop (tpsK0 zs m m) (esLoop_break (length zs) m) (tpsK_break zs m)" unfolding tmLoop_def esLoop_break_def proof (rule tm_loop_sem_false_traces[OF tmK1_break]) show "proper_symbols zs" using assms . show "read (tpsK_break zs m) ! 1 \ \" using tpsK_break_def read_def by simp show "esK_break (length zs) m @ [(1, m + 1)] = esK_break (length zs) m @ [(tpsK_break zs m :#: 0, tpsK_break zs m :#: 1)]" by (simp add: esK_break_def tpsK_break_def) qed definition esLoop :: "nat \ nat \ (nat \ nat) list" where "esLoop n m \ esLoop_done n m @ esLoop_break n m" lemma length_esLoop: "length (esLoop n m) \ m * (14 + 3 * m + (2 * m + 3) * n) + m + 2" unfolding esLoop_def using length_esLoop_done by (simp add: length_esLoop_break) lemma length_esLoop': "length (esLoop n m) \ 2 + 18 * m * m + 5 * m * m * n" proof - have "length (esLoop n m) \ m * (14 + 3 * m + (2 * m + 3) * n) + m + 2" using length_esLoop . also have "... = 14 * m + 3 * m * m + (2 * m * m + 3 * m) * n + m + 2" by algebra also have "... \ 15 * m * m + 3 * m * m + (2 * m * m + 3 * m) * n + 2" by simp also have "... \ 2 + 18 * m * m + 5 * m * m * n" by simp finally show ?thesis . qed lemma tmLoop: assumes "proper_symbols zs" shows "traces tmLoop (tpsK0 zs m 0) (esLoop (length zs) m) (tpsK_break zs m)" unfolding esLoop_def using assms by (intro traces_additive[OF tmLoop_done tmLoop_break]) lemma tmLoop': assumes "proper_symbols zs" shows "traces tmLoop (tpsB3 zs m) (esLoop (length zs) m) (tpsK_break zs m)" using assms tmLoop tpsK0_eq_tpsB3 by simp definition esB4 :: "nat \ nat \ (nat \ nat) list" where "esB4 n m \ esB3 n m @ esLoop n m" lemma length_esB4: "length (esB4 n m) \ 7 + 20 * m * m + 5 * m * m * n" proof - have "length (esB4 n m) \ 2 * m + 5 + m * (14 + 3 * m + (2 * m + 3) * n) + m + 2" unfolding esB4_def using length_esB3 length_esLoop - by (smt ab_semigroup_add_class.add_ac(1) add_less_cancel_left le_eq_less_or_eq length_append) + by (smt (verit) ab_semigroup_add_class.add_ac(1) add_less_cancel_left le_eq_less_or_eq length_append) also have "... = 2 * m + 5 + (14 * m + 3 * m * m + (2 * m * m + 3 * m) * n) + m + 2" by algebra also have "... = 7 + 17 * m + 3 * m * m + (2 * m * m + 3 * m) * n" by simp also have "... \ 7 + 17 * m + 3 * m * m + 5 * m * m * n" by simp also have "... \ 7 + 20 * m * m + 5 * m * m * n" by simp finally show ?thesis . qed lemma tmB4: assumes "proper_symbols zs" shows "traces tmB4 (tps_ones zs m) (esB4 (length zs) m) (tpsK_break zs m)" unfolding tmB4_def esB4_def using assms by (intro traces_sequential[OF tmB3 tmLoop']) definition tpsB5 :: "symbol list \ nat \ tape list" where "tpsB5 zs m \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m then \ else if x \ m + 1 + m * length zs then \ else 0, m + 1)]" definition esB5 :: "nat \ nat \ (nat \ nat) list" where "esB5 n m \ esB4 n m @ [(1, m + 1)]" lemma length_esB5: "length (esB5 n m) \ 8 + 20 * m * m + 5 * m * m * n" unfolding esB5_def using length_esB4 by (metis Suc_le_mono length_append_singleton one_plus_numeral plus_1_eq_Suc plus_nat.simps(2) semiring_norm(2) semiring_norm(4) semiring_norm(8)) lemma tmB5: assumes "proper_symbols zs" shows "traces tmB5 (tps_ones zs m) (esB5 (length zs) m) (tpsB5 zs m)" unfolding tmB5_def esB5_def proof (rule traces_sequential[OF tmB4]) show "proper_symbols zs" using assms . show "traces (tm_write 1 \) (tpsK_break zs m) [(1, m + 1)] (tpsB5 zs m)" proof (rule traces_tm_write_1I) show "1 < length (tpsK_break zs m)" using tpsK_break_def by simp_all show "[(1, m + 1)] = [(tpsK_break zs m :#: 0, tpsK_break zs m :#: 1)]" using tpsK_break_def by simp show "tpsB5 zs m = (tpsK_break zs m)[1 := tpsK_break zs m ! 1 |:=| \]" by (auto simp add: tpsK_break_def tpsB5_def) qed qed definition tpsB6 :: "symbol list \ nat \ tape list" where "tpsB6 zs m \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m then \ else if x \ m + 1 + m * length zs then \ else 0, 1)]" definition esB56 :: "nat \ nat \ (nat \ nat) list" where "esB56 n m \ map (Pair 1) (rev [0.. nat \ (nat \ nat) list" where "esB6 n m \ esB5 n m @ esB56 n m" lemma length_esB6: "length (esB6 n m) \ 11 + 21 * m * m + 5 * m * m * n" proof - have "length (esB6 n m) \ 8 + 20 * m * m + 5 * m * m * n + m + 3" unfolding esB6_def esB56_def using length_esB5 by (simp add: ab_semigroup_add_class.add_ac(1)) also have "... = 11 + 20 * m * m + 5 * m * m * n + m" by simp also have "... \ 11 + 21 * m * m + 5 * m * m * n" by simp finally show ?thesis . qed lemma tmB6: assumes "proper_symbols zs" shows "traces tmB6 (tps_ones zs m) (esB6 (length zs) m) (tpsB6 zs m)" unfolding tmB6_def esB6_def proof (rule traces_sequential[OF tmB5]) show "proper_symbols zs" using assms . show "traces (tm_cr 1) (tpsB5 zs m) (esB56 (length zs) m) (tpsB6 zs m)" proof (rule traces_tm_cr_1I) show "1 < length (tpsB5 zs m)" using tpsB5_def by simp show "clean_tape (tpsB5 zs m ! 1)" using tpsB5_def clean_tape_def by simp show "esB56 (length zs) m = map (Pair (tpsB5 zs m :#: 0)) (rev [0.. nat \ tape list" where "tpsB7 zs m \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m + 1 + m * length zs then \ else 0, m + 1)]" definition esB67 :: "nat \ nat \ (nat \ nat) list" where "esB67 n m \ map (\i. (1, 1 + Suc i)) [0.. nat \ (nat \ nat) list" where "esB7 n m \ esB6 n m @ esB67 n m" lemma length_esB7: "length (esB7 n m) \ 12 + 22 * m * m + 5 * m * m * n" proof - have "length (esB7 n m) \ 11 + 21 * m * m + 5 * m * m * n + m + 1" unfolding esB7_def esB67_def using length_esB6 - by (smt add.commute add_Suc_right add_le_cancel_right length_append length_append_singleton + by (smt (verit) add.commute add_Suc_right add_le_cancel_right length_append length_append_singleton length_map length_upt minus_nat.diff_0 plus_1_eq_Suc) also have "... \ 12 + 22 * m * m + 5 * m * m * n" by simp finally show ?thesis . qed lemma tmB7: assumes "proper_symbols zs" shows "traces tmB7 (tps_ones zs m) (esB7 (length zs) m) (tpsB7 zs m)" unfolding tmB7_def esB7_def proof (rule traces_sequential[OF tmB6]) show "proper_symbols zs" using assms . show "traces (tm_const_until 1 1 {\} \) (tpsB6 zs m) (esB67 (length zs) m) (tpsB7 zs m)" proof (rule traces_tm_const_until_11I) show "1 < length (tpsB6 zs m)" "3 < G" using tpsB6_def G G_ge4 by simp_all show "rneigh (tpsB6 zs m ! 1) {\} m" using tpsB6_def by (intro rneighI) auto show "esB67 (length zs) m = map (\i. (tpsB6 zs m :#: 0, tpsB6 zs m :#: 1 + Suc i)) [0.. m]" using constplant by (auto simp add: tpsB7_def tpsB6_def) qed qed definition tpsB8 :: "symbol list \ nat \ tape list" where "tpsB8 zs m \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m + 1 + m * length zs then \ else 0, 1)]" definition esB78 :: "nat \ nat \ (nat \ nat) list" where "esB78 n m \ map (Pair 1) (rev [0.. nat \ (nat \ nat) list" where "esB8 n m \ esB7 n m @ esB78 n m" lemma length_esB8: "length (esB8 n m) \ 15 + 23 * m * m + 5 * m * m * n" proof - have "length (esB8 n m) \ 12 + 22 * m * m + 5 * m * m * n + m + 3" unfolding esB8_def esB78_def using length_esB7 by (simp add: ab_semigroup_add_class.add_ac(1)) also have "... \ 15 + 23 * m * m + 5 * m * m * n" by simp finally show ?thesis . qed lemma tmB8: assumes "proper_symbols zs" shows "traces tmB8 (tps_ones zs m) (esB8 (length zs) m) (tpsB8 zs m)" unfolding tmB8_def esB8_def proof (rule traces_sequential[OF tmB7]) show "proper_symbols zs" using assms . show "traces (tm_cr 1) (tpsB7 zs m) (esB78 (length zs) m) (tpsB8 zs m)" proof (rule traces_tm_cr_1I) show "1 < length (tpsB7 zs m)" using tpsB7_def by simp show "clean_tape (tpsB7 zs m ! 1)" using tpsB7_def clean_tape_def by simp show "esB78 (length zs) m = map (Pair (tpsB7 zs m :#: 0)) (rev [0..Appending a fixed number of symbols\ text \ The next Turing machine appends a constant number $c$ of \textbf{1} symbols to the output tape. \ definition tmC :: "nat \ machine" where "tmC c \ tm_right_until 1 {\} ;; tm_write_repeat 1 \ c ;; tm_cr 1" lemma tmC_tm: "turing_machine 2 G (tmC c)" unfolding tmC_def using tm_right_until_tm tm_write_repeat_tm tm_cr_tm G by simp definition "tmC1 \ tm_right_until 1 {\}" definition "tmC2 c \ tmC1 ;; tm_write_repeat 1 \ c" definition "tmC3 c \ tmC2 c ;; tm_cr 1" definition tpsC1 :: "symbol list \ nat \ tape list" where "tpsC1 zs m \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m then \ else 0, m + 1)]" definition esC1 :: "nat \ nat \ (nat \ nat) list" where "esC1 n m \ map (\i. (1, 1 + Suc i)) [0..i. (tps5 zs m :#: 0, tps5 zs m :#: 1 + Suc i)) [0.. nat \ nat \ tape list" where "tpsC2 zs m c \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m + c then \ else 0, m + 1 + c)]" definition esC12 :: "nat \ nat \ nat \ (nat \ nat) list" where "esC12 n m c \ map (\i. (1, m + 1 + Suc i)) [0.. nat \ nat \ (nat \ nat) list" where "esC2 n m c \ esC1 n m @ esC12 n m c" lemma length_esC2: "length (esC2 n m c) = m + 1 + c" unfolding esC2_def by (simp add: length_esC1 esC12_def) lemma tmC2: assumes "proper_symbols zs" shows "traces (tmC2 c) (tps5 zs m) (esC2 (length zs) m c) (tpsC2 zs m c)" unfolding tmC2_def esC2_def proof (rule traces_sequential[OF tmC1]) show "proper_symbols zs" using assms . show "traces (tm_write_repeat 1 \ c) (tpsC1 zs m) (esC12 (length zs) m c) (tpsC2 zs m c)" proof (rule traces_tm_write_repeat_1I) show "1 < length (tpsC1 zs m)" using tpsC1_def by simp show "esC12 (length zs) m c = map (\i. (tpsC1 zs m :#: 0, tpsC1 zs m :#: 1 + Suc i)) [0.. c]" by (auto simp add: tpsC2_def tpsC1_def overwrite_def) qed qed definition tpsC3 :: "symbol list \ nat \ nat \ tape list" where "tpsC3 zs m c \ [(\zs\, 1), (\x. if x = 0 then \ else if x \ m + c then \ else 0, 1)]" definition esC23 :: "nat \ nat \ nat \ (nat \ nat) list" where "esC23 n m c \ map (Pair 1) (rev [0.. nat \ nat \ (nat \ nat) list" where "esC3 n m c \ esC2 n m c @ esC23 n m c" lemma length_esC3: "length (esC3 n m c) = 2 * m + 2 * c + 4" unfolding esC3_def by (simp add: length_esC2 esC23_def) lemma tmC3: assumes "proper_symbols zs" shows "traces (tmC3 c) (tps5 zs m) (esC3 (length zs) m c) (tpsC3 zs m c)" unfolding tmC3_def esC3_def proof (rule traces_sequential[OF tmC2]) show "proper_symbols zs" using assms . show "traces (tm_cr 1) (tpsC2 zs m c) (esC23 (length zs) m c) (tpsC3 zs m c)" proof (rule traces_tm_cr_1I) show "1 < length (tpsC2 zs m c)" using tpsC2_def by simp show "clean_tape (tpsC2 zs m c ! 1)" using tpsC2_def clean_tape_def by simp show "esC23 (length zs) m c = map (Pair (tpsC2 zs m c :#: 0)) (rev [0..Polynomials of higher degree\ text \ In order to construct polynomials of arbitrary degree, we repeat the TM @{term tmB}. \ fun tm_degree :: "nat \ machine" where "tm_degree 0 = []" | "tm_degree (Suc d) = tm_degree d ;; tmB" lemma tm_degree_tm: "turing_machine 2 G (tm_degree d)" proof (induction d) case 0 then show ?case by (simp add: turing_machine_def) next case (Suc d) then show ?case using turing_machine_def tmB_tm by (metis tm_degree.simps(2) turing_machine_sequential_turing_machine) qed text \ The number of \textbf{1} symbols the TM @{term "tm_degree d"} outputs on an input of length $n$: \ fun m_degree :: "nat \ nat \ nat" where "m_degree n 0 = n" | "m_degree n (Suc d) = 1 + m_degree n d * (n + 1)" fun es_degree :: "nat \ nat \ (nat * nat) list" where "es_degree n 0 = []" | "es_degree n (Suc d) = es_degree n d @ esB8 n (m_degree n d)" lemma tm_degree: assumes "proper_symbols zs" shows "traces (tm_degree d) (tps_ones zs (length zs)) (es_degree (length zs) d) (tps_ones zs (m_degree (length zs) d))" proof (induction d) case 0 then show ?case by fastforce next case (Suc d) have "traces (tm_degree d ;; tmB) (tps_ones zs (length zs)) (es_degree (length zs) d @ esB8 (length zs) (m_degree (length zs) d)) (tps_ones zs (m_degree (length zs) (Suc d)))" proof (rule traces_sequential[OF Suc.IH]) show "traces tmB (tps_ones zs (m_degree (length zs) d)) (esB8 (length zs) (m_degree (length zs) d)) (tps_ones zs (m_degree (length zs) (Suc d)))" using tmB[OF assms, of "m_degree (length zs) d"] m_degree.simps(2) by presburger qed then show ?case by simp qed text \ A lower bound for the number of \textbf{1} symbols the TM @{term "tm_degree d"} outputs: \ lemma m_degree_ge_pow: "m_degree n d \ n ^ (Suc d)" proof (induction d) case 0 then show ?case by simp next case (Suc d) have "m_degree n (Suc d) = 1 + m_degree n d * (n + 1)" by simp then have "m_degree n (Suc d) \ 1 + n ^ Suc d * (n + 1)" using Suc by (simp add: add_mono_thms_linordered_semiring(1)) then have "m_degree n (Suc d) \ 1 + n ^ Suc d * n + n ^ Suc d" by simp then have "m_degree n (Suc d) \ 1 + n ^ (Suc (Suc d)) + n ^ Suc d" by (metis power_Suc2) then show ?case by simp qed text \ An upper bound for the number of \textbf{1} symbols the TM @{term "tm_degree d"} outputs: \ lemma m_degree_poly: "big_oh_poly (\n. m_degree n d)" proof (induction d) case 0 have "(\n. m_degree n 0) = (\n. n)" by simp then show ?case using big_oh_poly_poly[of 1] by simp next case (Suc d) have "big_oh_poly (\n. n + 1)" using big_oh_poly_sum[OF big_oh_poly_poly[of 1] big_oh_poly_const[of 1]] by simp then have "big_oh_poly (\n. m_degree n d * (n + 1))" using big_oh_poly_prod[OF Suc] by blast then have "big_oh_poly (\n. 1 + m_degree n d * (n + 1))" using big_oh_poly_sum[OF big_oh_poly_const[of 1]] by simp then show ?case by simp qed corollary m_degree_plus_const_poly: "big_oh_poly (\n. m_degree n d + c)" using m_degree_poly big_oh_poly_sum big_oh_poly_const by simp lemma length_es_degree: "big_oh_poly (\n. length (es_degree n d))" proof (induction d) case 0 then show ?case using big_oh_poly_const by simp next case (Suc d) have "big_oh_poly (\n. length (esB8 n (m_degree n d)))" proof - let ?m = "\n. m_degree n d" have "big_oh_poly ?m" using m_degree_poly by simp then have "big_oh_poly (\n. 15 + 23 * ?m n * ?m n + 5 * ?m n * ?m n * n)" using big_oh_poly_sum big_oh_poly_const big_oh_poly_prod big_oh_poly_poly[of 1] by simp then show ?thesis using length_esB8 big_oh_poly_le by simp qed then show ?case using Suc big_oh_poly_sum by simp qed text \ Putting together the TM @{const tmA}, the TM @{term "tm_degree d"} for some $d$, and the TM @{term "tmC c"} for some $c$, we get a family of TMs parameterized by $d$ and $c$. These TMs construct all the polynomials we need. \ definition tm_poly :: "nat \ nat \ machine" where "tm_poly d c \ tmA ;; (tm_degree d ;; tmC c)" lemma tm_poly_tm: "turing_machine 2 G (tm_poly d c)" unfolding tm_poly_def using tmA_tm tm_degree_tm tmC_tm by simp definition es_poly :: "nat \ nat \ nat \ (nat \ nat) list" where "es_poly n d c \ es5 n @ es_degree n d @ esC3 n (m_degree n d) c" text \ On an input of length $n$ the Turing machine @{term "tm_poly d c"} outputs @{term "m_degree n d + c"} symbols~@{text \}. \ lemma tm_poly: assumes "proper_symbols zs" shows "traces (tm_poly d c) (tps0 zs) (es_poly (length zs) d c) (tps_ones zs (m_degree (length zs) d + c))" unfolding tm_poly_def es_poly_def using assms traces_sequential[OF tmA] traces_sequential[OF tm_degree] tmC by simp text \ The Turing machines run in polynomial time because their traces have polynomial length: \ lemma length_es_poly: "big_oh_poly (\n. length (es_poly n d c))" proof - have "big_oh_poly (\n. length (es5 n))" using length_es5 big_oh_poly_const big_oh_poly_prod big_oh_poly_sum big_oh_poly_poly[of 1] by simp moreover have "big_oh_poly (\n. length (esC3 n (m_degree n d) c))" proof - have *: "(\n. length (esC3 n (m_degree n d) c)) = (\n. 2 * (m_degree n d) + 2 * c + 4)" using length_esC3 by fast have "big_oh_poly (\n. 2 * (m_degree n d) + 2 * c + 4)" using m_degree_poly big_oh_poly_const big_oh_poly_prod big_oh_poly_sum by simp then show ?thesis by (simp add: *) qed ultimately have "big_oh_poly (\n. length (es5 n) + length (es_degree n d) + length (esC3 n (m_degree n d) c))" using length_es_degree big_oh_poly_sum by blast then show ?thesis by (simp add: es_poly_def add.assoc) qed text \ The Turing machine @{term "tm_poly d c"} outputs @{term "m_degree n c + c"} many @{text \} symbols on an input of length $n$. Hence for every polynomially bounded function $f$ there is such a Turing machine outputting at least $f(n)$ symbols @{text \}. \ lemma m_degree_plus_const: assumes "big_oh_poly f" obtains d c where "\n. f n \ m_degree n d + c" proof - obtain c m d where f: "\n>m. f n \ c * n ^ d" using assms big_oh_poly by auto let ?d = "Suc d" let ?k = "max c m" have "n ^ ?d \ c * n ^ d" if "n > ?k" for n using that by simp moreover have "f n \ c * n ^ d" if "n > ?k" for n using f that by simp ultimately have 1: "f n \ n ^ ?d" if "n > ?k" for n using that using order_trans by blast define c' where "c' = Max {f n| n. n \ ?k}" moreover have "finite {f n| n. n \ ?k}" by simp ultimately have "c' \ f n" if "n \ ?k" for n using that Max.bounded_iff by blast then have "f n \ n ^ ?d + c'" if "n \ ?k" for n by (simp add: that trans_le_add2) moreover have "f n \ n ^ ?d + c'" if "n > ?k" for n using that 1 by fastforce ultimately have "f n \ n ^ ?d + c'" for n using leI by blast then have "f n \ m_degree n d + c'" for n using m_degree_ge_pow by (meson le_diff_conv less_le_trans not_le) then show ?thesis using that by auto qed text \ The Turing machines are oblivious. \ lemma tm_poly_oblivious: "oblivious (tm_poly d c)" proof - have tm: "turing_machine 2 G (tm_poly d c)" using tm_poly_tm by simp have init: "(0, tps0 zs) = start_config 2 zs" for zs using tps0_def start_config_def contents_def by auto { fix zs assume "bit_symbols zs" then have proper: "proper_symbols zs" by auto define tps where "tps = tps_ones zs (m_degree (length zs) d + c)" moreover define e where "e = (\n. es_poly n d c)" ultimately have "trace (tm_poly d c) (start_config 2 zs) (e (length zs)) (length (tm_poly d c), tps)" using tm_poly init proper by (simp add: traces_def) } then show ?thesis using tm oblivious_def by fast qed end (* locale turing_machine_poly *) definition start_tapes_2 :: "symbol list \ tape list" where "start_tapes_2 zs \ [(\zs\, 0), (\i. if i = 0 then \ else \, 0)]" definition one_tapes_2 :: "symbol list \ nat \ tape list" where "one_tapes_2 zs m \ [(\zs\, 1), (\replicate m \\, 1)]" text \ The main result of this chapter. For every polynomially bounded function $g$ there is a polynomial-time two-tape oblivious Turing machine that outputs at least $g(n)$ symbols~@{text \} for every input length $n$. \ lemma polystructor: assumes "big_oh_poly g" and "G \ 5" shows "\M es f. turing_machine 2 G M \ big_oh_poly (\n. length (es n)) \ big_oh_poly f \ (\n. g n \ f n) \ (\zs. proper_symbols zs \ traces M (start_tapes_2 zs) (es (length zs)) (one_tapes_2 zs (f (length zs))))" proof - interpret turing_machine_poly G using assms(2) by (simp add: turing_machine_poly_def) obtain d c where dc: "\n. g n \ m_degree n d + c" using assms(1) m_degree_plus_const by auto define M where "M = tm_poly d c" define es where "es = (\n. es_poly n d c)" define f where "f = (\n. m_degree n d + c)" have "turing_machine 2 G M" using M_def tm_poly_tm assms(2) by simp moreover have "big_oh_poly (\n. length (es n))" using es_def length_es_poly by simp moreover have "\n. g n \ f n" using f_def dc by simp moreover have "big_oh_poly f" using f_def m_degree_plus_const_poly by simp moreover have "\zs. proper_symbols zs \ traces M (start_tapes_2 zs) (es (length zs)) (one_tapes_2 zs (f (length zs)))" proof (standard, standard) fix zs assume zs: "proper_symbols zs" have "traces M (tps0 zs) (es (length zs)) (tps_ones zs (f (length zs)))" unfolding M_def es_def f_def using tm_poly[OF zs, of d c] by simp moreover have "tps0 = start_tapes_2" using tps0_def start_tapes_2_def by presburger ultimately have "traces M (start_tapes_2 zs) (es (length zs)) (tps_ones zs (f (length zs)))" by simp moreover have "one_tapes_2 = tps_ones" using one_tapes_2_def contents_def by fastforce ultimately show "traces M (start_tapes_2 zs) (es (length zs)) (one_tapes_2 zs (f (length zs)))" by metis qed ultimately show ?thesis by auto qed end \ No newline at end of file diff --git a/thys/Cook_Levin/Reducing.thy b/thys/Cook_Levin/Reducing.thy --- a/thys/Cook_Levin/Reducing.thy +++ b/thys/Cook_Levin/Reducing.thy @@ -1,3276 +1,3277 @@ chapter \Reducing $\NP$ languages to \SAT{}\label{s:Reducing}\ theory Reducing imports Satisfiability Oblivious begin text \ We have already shown that \SAT{} is in $\NP$. It remains to show that \SAT{} is $\NP$-hard, that is, that every language $L \in \NP$ can be polynomial-time reduced to \SAT{}. This, in turn, can be split in two parts. First, showing that for every $x$ there is a CNF formula $\Phi$ such that $x\in L$ iff.\ $\Phi$ is satisfiable. Second, that $\Phi$ can be computed from $x$ in polynomial time. This chapter is devoted to the first part, which is the core of the proof. In the subsequent two chapters we painstakingly construct a polynomial-time Turing machine computing $\Phi$ from $x$ in order to show something that is usually considered ``obvious''. The proof corresponds to lemma~2.11 from the textbook~\cite{ccama}. Of course we have to be much more explicit than the textbook, and the first section describes in some detail how we derive the formula $\Phi$. \ section \Introduction\ text \ Let $L \in \NP$. In order to reduce $L$ to \SAT{}, we need to construct for every string $x\in\bbOI^*$ a CNF formula $\Phi$ such that $x\in L$ iff.\ $\Phi$ is satisfiable. In this section we describe how $\Phi$ looks like. \subsection{Preliminaries} We denote the length of a string $s\in\bbOI^*$ by $|s|$. We define \[ num(s) = \left\{\begin{array}{ll} k & \text{if } s = \bbbI^k\bbbO^{|s|-k},\\ |s| + 1 & \text{otherwise.} \end{array}\right. \] Essentially $num$ interprets some strings as unary codes of numbers. All other strings are interpreted as an ``error value''. For a string $s$ and a sequence $w\in\nat^n$ of numbers we write $s(w)$ for $num(s_{w_0}\dots s_{w_{n-1}})$. Likewise for an assignment $\alpha\colon \nat \to \bbOI$ we write $\alpha(w) = num(\alpha(w_0)\dots\alpha(w_{n-1}))$. We define two families of CNF formulas. Variables are written $v_0, v_1, v_2, \dots$, and negated variables are written $\bar v_0, \bar v_1, \bar v_2, \dots$ Let $w\in\nat^n$ be a list of numbers. For $k \leq n$ define \[ \Psi(w, k) = \bigwedge_{i=0}^{k-1} v_{w_i} \land \bigwedge_{i=k}^{n - 1} \bar v_{w_i}. \] This formula is satisfied by an assignment $\alpha$ iff.\ $\alpha(w) = k$. In a similar fashion we define for $n > 2$, \[ \Upsilon(w) = v_{w_0} \land v_{w_1} \land \bigwedge_{i=3}^{n - 1} \bar v_{w_i}, \] which is satisfied by an assignment $\alpha$ iff.\ $\alpha(w) \in \{2, 3\} = \{\mathbf{0}, \mathbf{1}\}$, where as usual the boldface $\mathbf{0}$ and $\mathbf{1}$ refer to the symbols represented by the numbers~2 and~3. For $a \leq b$ we write $[a:b]$ for the interval $[a, \dots, b - 1] \in \nat^{b - a}$. For intervals the CNF formulas become: \[ \Psi([a:b], k) = \bigwedge_{i=a}^{a+k-1} v_i \land \bigwedge_{i=a+k}^{b - 1} \bar v_i \qquad \text{and} \qquad \Upsilon([a:b]) = v_a \land v_{a+1} \land \bigwedge_{i=a+3}^{b - 1} \bar v_i. \] Let $\phi$ be a CNF formula and let $\sigma\in\nat^*$ be a sequence of variable indices such that for all variables $v_i$ occurring in $\phi$ we have $i < |\sigma|$. Then we define the CNF formula $\sigma(\phi)$ as the formula resulting from replacing every variable $v_i$ in $\phi$ by the variable $v_{\sigma_i}$. This corresponds to our function @{const relabel}. \subsection{Construction of the CNF formula} Let $M$ be the two-tape oblivious verifier Turing machine for $L$ from lemma~@{text NP_imp_oblivious_2tape}. Let $p$ be the polynomial function for the length of the certificates, and let $T\colon \nat \to \nat$ be the polynomial running-time bound. Let $G$ be $M$'s alphabet size. Let $x\in\bbOI^n$ be fixed thoughout the rest of this section. We seek a CNF formula $\Phi$ that is satisfiable iff.\ $x\in L$. We are going to transform ``$x\in L$'' via several equivalent statements to the statement ``$\Phi$ is satisfiable'' for a suitable $\Phi$ defined along the way. The Isabelle formalization later in this chapter does not prove these equivalences explicitly. They are only meant to explain the shape of $\Phi$. \subsubsection{1st equivalence} From lemma~@{text NP_imp_oblivious_2tape} about $M$ we get the first equivalent statement: There exists a certificate $u \in \bbOI^{p(n)}$ such that $M$ on input $\langle x, u\rangle$ halts with the symbol \textbf{1} under its output tape head. The running time of $M$ is bounded by $T(|\langle x, u\rangle|) = T(2n+2+2p(n))$. We abbreviate $|\langle x, u\rangle| = 2n+2+2p(n)$ by $m$. \subsubsection{2nd equivalence} For the second equivalent statement, we define what the textbook calls ``snapshots''. For every $u\in\{\bbbO,\bbbI\}^{p(n)}$ let $z_0^u(t)$ be the symbol under the input tape head of $M$ on input $\langle x, u\rangle$ at step $t$. Similarly we define $z_1^u(t)$ as the symbol under the output tape head of $M$ at step $t$ and $z_2^u(t)$ as the state $M$ is in at step $t$. A triple $z^u(t) = (z^u_0(t), z^u_1(t), z^u_2(t))$ is called a snapshot. For the initial snapshot we have: \begin{equation} z_0^u(0) = z_1^u(0) = \triangleright\qquad\text{and}\qquad z_2^u(0) = 0. \tag{Z0} \end{equation} The crucial idea is that the snapshots for $t > 0$ can be characterized recursively using two auxiliary functions $\inputpos$ and $\prev$. Since $M$ is oblivious, the positions of the tape heads on input $\langle x, u\rangle$ after $t$ steps are the same for all $u$ of length $p(n)$. We denote the input head positions by $\inputpos(t)$. For every $t$ we denote by $\prev(t)$ the last step before $t$ in which the output tape head of $M$ was in the same cell as in step $t$. Due to $M$'s obliviousness this is again the same for all $u$ of length $p(n)$. If there is no such previous step, because $t$ is the first time the cell is reached, we set $\prev(t) = t$. (This deviates from the textbook, which sets $\prev(t) = 1$.) In the other case we have $\prev(t) < t$. Also due to $M$'s obliviousness, the halting time on input $\langle x, u\rangle$ is the same for all $u$ of length $p(n)$, and we denote it by $T' \le T(|\langle x, u\rangle|)$. Thus we have $\inputpos(t) \le T'$ for all $t$. If we define the symbol sequence $y(u) = \triangleright \langle x, u\rangle \Box^{T'}$, the first component of the snapshots is, for arbitrary $t$: \begin{equation} z_0^u(t) = y(u)_{\inputpos(t)}. \tag{Z1} \end{equation} Next we consider the snapshot components $z_1^u(t)$ for $t > 0$. First consider the case $\prev(t) < t$; that is, the last time before $t$ when $M$'s output tape head was in the same cell as in step $t$ was in step $\prev(t)$. The snapshot for step $\prev(t)$ has exactly the information needed to calculate the actions of $M$ at step $t$: the symbols read from both tapes and the state which $M$ is in. In some sort of hybrid notation: \begin{equation} z_1^u(t) = (M\ !\ z_2^u(\prev(t)))\ [z_0^u(\prev(t)), z_1^u(\prev(t))]\ [.]\ 1. \tag{Z2} \end{equation} In the other case, $\prev(t) = t$, the output tape head has not been in this cell before and is thus reading a blank. It cannot be reading the start symbol because the output tape head was in cell zero at step $t = 0$ already. Formally: \begin{equation} z_1^u(t) = \Box. \tag{Z3} \end{equation} The state $z_2^u(t)$ for $t > 0$ can be computed from the state $z_2^u(t-1)$ in the previous step and the symbols $z_0^u(t-1)$ and $z_1^u(t-1)$ read in the previous step: \begin{equation} z_2^u(t) = \mathit{fst}\ ((M\ !\ z_2^u(t - 1))\ [z_0^u(t - 1), z_1^u(t - 1)]). \tag{Z4} \end{equation} For a string $u\in\bbOI^{p(n)}$ the equations (Z0) -- (Z4) uniquely determine all the $z^u(0), \dots, z^u(T')$. Conversely, the snapshots for $u$ satisfy all the equations. Therefore the equations characterize the sequence of snapshots. The condition that $M$ halts with the output tape head on \textbf{1} can be expressed with snapshots: \begin{equation} z_1^{u}(T') = \mathbf{1}. \tag{Z5} \end{equation} This yields our second equivalent statement: $x\in L$ iff.\ there is a $u\in\{\bbbO,\bbbI\}^{p(n)}$ and a sequence $z^u(0), \dots, z^u(T')$ satisfying the equations (Z0) -- (Z5). \subsubsection{3rd equivalence} The length of $y(u)$ is $m' := m + 1 + T' = 3 + 2n + 2p(n) + T'$ because $|\langle x, u\rangle| = m$ plus the start symbol plus the $T'$ blanks. For the next equivalence we observe that the strings $y(u)$ for $u\in\{\bbbO,\bbbI\}^{p(n)}$ can be characterized as follows. Consider a predicate on strings $y$: \begin{itemize} \item[(Y0)] $|y| = m'$; \item[(Y1)] $y_0 = \triangleright$ (the start symbol); \item[(Y2)] $y_{2n+1} = y_{2n+2} = \mathbf{1}$ (the separator in the pair encoding); \item[(Y3)] $y_{2i+1} = \mathbf{0}$ for $i=0,\dots,n-1$ (the zeros before $x$); \item[(Y4)] $y_{2n+2+2i+1} = \mathbf{0}$ for $i=0, \dots, p(n)-1$ (the zeros before $u$); \item[(Y5)] $y_{2n+2p(n)+3+i} = \Box$ for $i=0, \dots, T' - 1$ (the blanks after the input proper); \item[(Y6)] $y_{2i+2} = \left\{\begin{array}{ll} \mathbf{0} & \text{if } x_i = \bbbO,\\ \mathbf{1} & \text{otherwise} \end{array}\right.$ for $i = 0, \dots, n - 1$ (the bits of $x$); \item[(Y7)] $y_{2n+4+2i} \in \{\mathbf{0}, \mathbf{1}\}$ for $i=0,\dots,p(n)-1$ (the bits of $u$). \end{itemize} Every $y(u)$ for some $u$ of length $p(n)$ satisfies this predicate. Conversely, from a $y$ satisfying the predicate, a $u$ of length $p(n)$ can be extracted such that $y = y(u)$. From that we get the third equivalent statement: $x \in L$ iff.\ there is a $y \in \{0, \dots, G - 1\}^{m'}$ with (Y0) -- (Y7) and a sequence $z^u(0), \dots, z^u(T')$ with (Z0) -- (Z5). \subsubsection{4th equivalence} Each element of $y$ is a symbol from $M$'s alphabet, that is, a number less than $G$. The same goes for the first two elements of each snapshot, $z_0^u(t)$ and $z_1^u(t)$. The third element, $z_2^u(t)$, is a number less than or equal to the number of states of $M$. Let $H$ be the maximum of $G$ and the number of states. Every element of $y$ and of the snapshots can then be represented by a bit string of length $H$ using $num$ (the textbook uses binary, but unary is simpler for us). So we use $3H$ bits to represent one snapshot. There are $T' + 1$ snapshots until $M$ halts. Thus all elements of all snapshots can be represented by a string of length $3H\cdot(T'+1)$. Together with the string of length $N := H\cdot m'$ for the input tape contents $y$, we have a total length of $N + 3H\cdot(T'+1)$. The equivalence can thus be stated as $x \in L$ iff.\ there is a string $s\in \{\bbbO,\bbbI\}^{N + 3H\cdot(T'+1)}$ with certain properties. To write these properties we introduce some intervals: \begin{itemize} \item $\gamma_i = [iH : (i+1)H]$ for $i < m'$, \item $\zeta_0(t) = [N + 3Ht : N+3Ht + H]$ for $t \leq T'$, \item $\zeta_1(t) = [N + 3Ht + H : N+3Ht + 2H]$ for $t \leq T'$, \item $\zeta_2(t) = [N + 3Ht + 2H : N+3H(t + 1)]$ for $t \leq T'$. \end{itemize} These intervals slice the string $s$ in intervals of length $H$. The string $s$ must satisfy properties analogous to (Y0) -- (Y7), which we express using the intervals $\gamma_i$: \begin{itemize} \item[(Y0)] $|s| = N + 3H(T' + 1)$ \item[(Y1)] $s(\gamma_0) = \triangleright$ (the start symbol); \item[(Y2)] $s(\gamma_{2n+1}) = s(\gamma_{2n+2}) = \mathbf{1}$ (the separator in the pair encoding); \item[(Y3)] $s(\gamma_{2i+1}) = \mathbf{0}$ for $i=0,\dots,n-1$ (the zeros before $x$); \item[(Y4)] $s(\gamma_{2n+2+2i+1}) = \mathbf{0}$ for $i=0,\dots,p(n)-1$ (the zeros before $u$); \item[(Y5)] $s(\gamma_{2n+2p(n)+3+i}) = \Box$ for $i=0,\dots,T' - 1$ (the blanks after the input proper); \item[(Y6)] $s(\gamma_{2i+2}) = \left\{\begin{array}{ll} \mathbf{0} & \text{if } x_i = \bbbO,\\ \mathbf{1} & \text{otherwise} \end{array}\right.$ for $i = 0, \dots, n - 1$ (the bits of $x$); \item[(Y7)] $s(\gamma_{2n+4+2i}) \in \{\mathbf{0}, \mathbf{1}\}$ for $i=0,\dots,p(n)-1$ (the bits of $u$). \end{itemize} Moreover the string $s$ must satisfy (Z0) -- (Z5). For these properties we use the $\zeta$ intervals. \begin{itemize} \item[(Z0)] $s(\zeta_0(0)) = s(\zeta_1(0)) = \triangleright$ and $s(\zeta_2(0)) = 0$, \item[(Z1)] $s(\zeta_0(t)) = s(\gamma_{\inputpos(t)})$ for $t = 1, \dots, T'$, \item[(Z2)] $s(\zeta_1(t)) = (M\ !\ s(\zeta_2(\prev(t)))\ [s(\zeta_0(\prev(t))), s(\zeta_1(\prev(t)))]\ [.]\ 1$ for $t = 1, \dots, T'$ with $\prev(t) < t$, \item[(Z3)] $s(\zeta_1(t)) = \Box$ for $t = 1, \dots, T'$ with $\prev(t) = t$, \item[(Z4)] $s(\zeta_2(t)) = \mathit{fst}\ ((M\ !\ s(\zeta_2(t - 1))\ [s(\zeta_0(t - 1), s(\zeta_1(t - 1)])$ for $t = 1, \dots, T'$, \item[(Z5)] $s(\zeta_1(T')) = \mathbf{1}$. \end{itemize} \subsubsection{5th equivalence} An assignment is an infinite bit string. For formulas over variables with indices in the interval $[0 : N+3H(T'+1)]$, only the initial $N+3H(T'+1)$ bits of the assignment are relevant. If we had a CNF formula $\Phi$ over these variables that is satisfied exactly by the assignments $\alpha$ for which there is an $s$ with the above properties and $\alpha(i) = s_i$ for all $i < |s|$, then the existence of such an $s$ would be equivalent to $\Phi$ being satisfiable. Next we construct such a CNF formula. Most properties are easy to translate using the formulas $\Psi$ and $\Upsilon$. For example, $s(\gamma_0) = \triangleright$ corresponds to $\alpha(\gamma_0) = \triangleright$. A formula that is satisfied exactly by assignments with this property is $\Psi(\gamma_0, 1)$. Likewise the property $s(\gamma_{2n+4+2i})\in\{\mathbf{0}, \mathbf{1}\}$ corresponds to the CNF formula $\Upsilon(\gamma_{2n+4+2i})$. Property (Y0) corresponds to $\Phi$ having only variables $0, \dots, N+3H(T'+1) - 1$. The other (Y$\cdot$) properties become: \begin{itemize} \item[(Y1)] $\Phi_1 := \Psi(\gamma_0, 1)$, \item[(Y2)] $\Phi_2 := \Psi(\gamma_{2n+1}, 3) \land \Psi(\gamma_{2n+2}, 3)$, \item[(Y3)] $\Phi_3 := \bigwedge_{i=0}^{n-1}\Psi(\gamma_{2i+1}, 2)$, \item[(Y4)] $\Phi_4 := \bigwedge_{i=0}^{p(n) - 1}\Psi(\gamma_{2n+2+2i+1}, 2)$, \item[(Y5)] $\Phi_5 := \bigwedge_{i=0}^{T' - 1}\Psi(\gamma_{2n+2p(n)+3+i}, 0)$, \item[(Y6)] $\Phi_6 := \bigwedge_{i=0}^{n - 1}\Psi(\gamma_{2i+2}, x_i + 2)$, \item[(Y7)] $\Phi_7 := \bigwedge_{i=0}^{p(n) - 1}\Upsilon(\gamma_{2n+4+2i})$. \end{itemize} Property~(Z0) and property~(Z5) become these formulas: \begin{itemize} \item[(Z0)] $\Phi_0 := \Psi(\zeta_0(0), 1) \land \Psi(\zeta_1(0), 1) \land \Psi(\zeta_2(0), 0)$, \item[(Z5)] $\Phi_8 := \Psi(\zeta_1(T'), 3)$. \end{itemize} The remaining properties (Z1) -- (Z4) are more complex. They apply to $t = 1, \dots T'$. Let us first consider the case $\prev(t) < t$. With $\alpha$ for $s$ the properties become: \begin{itemize} \item[(Z1)] $\alpha(\zeta_0(t)) = \alpha(\gamma_{\inputpos(t)})$, \item[(Z2)] $\alpha(\zeta_1(t)) = ((M\ !\ \alpha(\zeta_2(\prev(t))))\ [\alpha(\zeta_0(\prev(t))), \alpha(\zeta_1(\prev(t)))])\ [.]\ 1$, \item[(Z4)] $\alpha(\zeta_2(t)) = \mathit{fst}\ ((M\ !\ \alpha(\zeta_2(t-1)))\ [\alpha(\zeta_0(t-1)), \alpha(\zeta_1(t-1))])$. \end{itemize} For any $t$ the properties (Z1), (Z2), (Z4) use at most $10H$ variable indices, namely all the variable indices in the nine $\zeta$'s and in $\gamma_{\inputpos(t)}$, each of which have $H$ indices. Now if the set of all these variable indices was $\{0, \dots, 10H - 1\}$ we could apply lemma~@{thm [source] depon_ex_formula} to get a CNF formula $\psi$ over these variables that ``captures the spirit'' of the properties. We would then merely have to relabel the formula with the actual variable indices we need for each $t$. More precisely, let $w_i = [iH : (i+1)H]$ for $i = 0, \dots, 9$ and consider the following criterion for $\alpha$ on the variable indices $\{0, \dots 10H - 1\}$: \begin{itemize} \item[($\mathrm{F}_1$)] $\alpha(w_6) = \alpha(w_9)$, \item[($\mathrm{F}_2$)] $\alpha(w_7) = ((M\ !\ \alpha(w_5))\ [\alpha(w_3), \alpha(w_4)])\ [.]\ 1,$ \item[($\mathrm{F}_3$)] $\alpha(w_8) = \mathit{fst}\ ((M\ !\ \alpha(w_2))\ [\alpha(w_0), \alpha(w_1)]).$ \end{itemize} Lemma~@{thm [source] depon_ex_formula} gives us a formula $\psi$ satisfied exactly by those assignments $\alpha$ that meet the conditions ($\mathrm{F}_1$), ($\mathrm{F}_2$), ($\mathrm{F}_3$). From this $\psi$ we can create all the formulas we need for representing the properties~(Z1), (Z2), (Z4) by substituting (``relabeling'' in our terminology) the variables $[0,10H)$ appropriately. The substitution for step $t$ is \[ \rho_t = \zeta_0(t-1) \circ \zeta_1(t-1) \circ \zeta_2(t-1) \circ \zeta_0(\prev(t)) \circ \zeta_1(\prev(t)) \circ \zeta_2(\prev(t)) \circ \zeta_0(t) \circ \zeta_1(t) \circ \zeta_2(t) \circ \gamma_{\inputpos(t)}, \] where $\circ$ denotes the concatenation of lists. Then $\rho_t(\psi)$ is CNF formula satisfied exactly by the assignments $\alpha$ satisfying (Z1), (Z2), (Z4). For the case $\prev(t) = t$ we have a criterion on the variable indices $\{0, \dots, 7H - 1\}$: \begin{itemize} \item[($\mathrm{F}_1'$)] $\alpha(w_3) = \alpha(w_6)$, \item[($\mathrm{F}_2'$)] $\alpha(w_4) = \Box$, \item[($\mathrm{F}_3'$)] $\alpha(w_5) = \mathit{fst}\ ((M\ !\ \alpha(w_2))\ [\alpha(w_0), \alpha(w_1)])$, \end{itemize} whence lemma~@{thm [source] depon_ex_formula} supplies us with a formula $\psi'$. With appropriate substitutions \[ \rho'_t = \zeta_0(t-1) \circ \zeta_1(t-1) \circ \zeta_2(t-1) \circ \zeta_0(t) \circ \zeta_1(t) \circ \zeta_2(t) \circ \gamma_{\inputpos(t)}, \] we then define CNF formulas $\chi_t$ for all $t = 1, \dots, T'$: \[ \chi_t = \left\{\begin{array}{ll} \rho_t(\psi) &\text{if }\prev(t) < t,\\ \rho'_t(\psi') &\text{if }\prev(t) = t. \end{array}\right. \] The point of all that is that we can hard-code $\psi$ and $\psi'$ in the TM performing the reduction (to be built in the final chapter) and for each $t$ the TM only needs to construct the substitution $\rho_t$ or $\rho_t'$ and perform the relabeling. Turing machines that perform these operations will be constructed in the next chapter. Since all $\chi_t$ are in CNF, so is the conjunction \[ \Phi_9 := \bigwedge_{t=1}^{T'}\chi_t\ . \] Finally the complete CNF formula is: \[ \Phi := \Phi_0 \land \Phi_1 \land \Phi_2 \land \Phi_3 \land \Phi_4 \land \Phi_5 \land \Phi_6 \land \Phi_7 \land \Phi_8 \land \Phi_9\ . \] \ section \Auxiliary CNF formulas\ text \ In this section we define the CNF formula families $\Psi$ and $\Upsilon$. In the introduction both families were parameterized by intervals of natural numbers. Here we generalize the definition to allow arbitrary sequences of numbers although we will not need this generalization. \ text \ The number of variables set to true in a list of variables: \ definition numtrue :: "assignment \ nat list \ nat" where "numtrue \ vs \ length (filter \ vs)" text \ Checking whether the list of bits assigned to a list @{term vs} of variables has the form $\bbbI\dots\bbbI\bbbO\dots\bbbO$: \ definition blocky :: "assignment \ nat list \ nat \ bool" where "blocky \ vs k \ \i (vs ! i) \ i < k" text \ The next function represents the notation $\alpha(\gamma)$ from the introduction, albeit generalized to lists that are not intervals $\gamma$. \ definition unary :: "assignment \ nat list \ nat" where "unary \ vs \ if (\k. blocky \ vs k) then numtrue \ vs else Suc (length vs)" lemma numtrue_remap: assumes "\s\set seq. s < length \" shows "numtrue (remap \ \) seq = numtrue \ (reseq \ seq)" proof - have *: "length (filter f xs) = length (filter (f \ ((!) xs)) [0.. ((!) seq)) [0.. ?s_seq = length (filter \ ?s_seq)" using numtrue_def by simp then have "numtrue \ ?s_seq = length (filter (\ \ ((!) ?s_seq)) [0.. ?s_seq = length (filter (\ \ ((!) ?s_seq)) [0.. ((!) seq)) i = (\ \ ((!) ?s_seq)) i" if "i < length seq" for i using assms remap_def reseq_def that by simp then show ?thesis using lhs rhs by (metis atLeastLessThan_iff filter_cong set_upt) qed lemma unary_remap: assumes "\s\set seq. s < length \" shows "unary (remap \ \) seq = unary \ (reseq \ seq)" proof - have *: "blocky (remap \ \) seq k = blocky \ (reseq \ seq) k" for k using blocky_def remap_def reseq_def assms by simp let ?alpha = "remap \ \" and ?seq = "reseq \ seq" show ?thesis proof (cases "\k. blocky ?alpha seq k") case True then show ?thesis using * unary_def numtrue_remap assms by simp next case False then have "unary ?alpha seq = Suc (length seq)" using unary_def by simp moreover have "\ (\k. blocky \ ?seq k)" using False * assms by simp ultimately show ?thesis using unary_def by simp qed qed text \ Now we define the family $\Psi$ of CNF formulas. It is parameterized by a list @{term vs} of variable indices and a number $k \leq |\mathit{vs}|$. The formula is satisfied exactly by those assignments that set to true the first $k$ variables in $vs$ and to false the other variables. This is more general than we need, because for us $vs$ will always be an interval. \ definition Psi :: "nat list \ nat \ formula" ("\") where "\ vs k \ map (\s. [Pos s]) (take k vs) @ map (\s. [Neg s]) (drop k vs)" lemma Psi_unary: assumes "k \ length vs" and "\ \ \ vs k" shows "unary \ vs = k" proof - have "\ \ map (\s. [Pos s]) (take k vs)" using satisfies_def assms Psi_def by simp then have "satisfies_clause \ [Pos v]" if "v \ set (take k vs)" for v using satisfies_def that by simp then have "satisfies_literal \ (Pos v)" if "v \ set (take k vs)" for v using satisfies_clause_def that by simp then have pos: "\ v" if "v \ set (take k vs)" for v using that by simp have "\ \ map (\s. [Neg s]) (drop k vs)" using satisfies_def assms Psi_def by simp then have "satisfies_clause \ [Neg v]" if "v \ set (drop k vs)" for v using satisfies_def that by simp then have "satisfies_literal \ (Neg v)" if "v \ set (drop k vs)" for v using satisfies_clause_def that by simp then have neg: "\ \ v" if "v \ set (drop k vs)" for v using that by simp have "blocky \ vs k" proof - have "\ (vs ! i)" if "i < k" for i using pos that assms(1) by (metis in_set_conv_nth length_take min_absorb2 nth_take) moreover have "\ \ (vs ! i)" if "i \ k" "i < length vs" for i using that assms(1) neg by (metis Cons_nth_drop_Suc list.set_intros(1) set_drop_subset_set_drop subsetD) ultimately show ?thesis using blocky_def by (metis linorder_not_le) qed moreover have "numtrue \ vs = k" proof - have "filter \ vs = take k vs" using pos neg by (metis (mono_tags, lifting) append.right_neutral append_take_drop_id filter_True filter_append filter_empty_conv) then show ?thesis using numtrue_def assms(1) by simp qed ultimately show ?thesis using unary_def by auto qed text \ We will only ever consider cases where $k \leq |vs|$. So we can use $blocky$ to show that an assignment satisfies a $\Psi$ formula. \ lemma satisfies_Psi: assumes "k \ length vs" and "blocky \ vs k" shows "\ \ \ vs k" proof - have "\ \ map (\s. [Pos s]) (take k vs)" (is "\ \ ?phi") proof - { fix c :: clause assume c: "c \ set ?phi" then obtain s where "c = [Pos s]" and "s \ set (take k vs)" by auto then obtain i where "i < k" and "s = vs ! i" using assms(1) by (smt (verit, best) in_set_conv_nth le_antisym le_trans nat_le_linear nat_less_le nth_take nth_take_lemma take_all_iff) then have "c = [Pos (vs ! i)]" using `c = [Pos s]` by simp moreover have "i < length vs" using assms(1) `i < k` by simp ultimately have "\ (vs ! i)" using assms(2) blocky_def \i < k\ by blast then have "satisfies_clause \ c" using satisfies_clause_def by (simp add: \c = [Pos (vs ! i)]\) } then show ?thesis using satisfies_def by simp qed moreover have "\ \ map (\s. [Neg s]) (drop k vs)" (is "\ \ ?phi") proof - { fix c :: clause assume c: "c \ set ?phi" then obtain s where "c = [Neg s]" and "s \ set (drop k vs)" by auto then obtain j where "j < length vs - k" and "s = drop k vs ! j" by (metis in_set_conv_nth length_drop) define i where "i = j + k" then have "i \ k" and "s = vs ! i" by (auto simp add: \s = drop k vs ! j\ add.commute assms(1) i_def) then have "c = [Neg (vs ! i)]" using `c = [Neg s]` by simp moreover have "i < length vs" using assms(1) using \j < length vs - k\ i_def by simp ultimately have "\ \ (vs ! i)" using assms(2) blocky_def[of \ vs k] i_def by simp then have "satisfies_clause \ c" using satisfies_clause_def by (simp add: \c = [Neg (vs ! i)]\) } then show ?thesis using satisfies_def by simp qed ultimately show ?thesis using satisfies_def Psi_def by auto qed lemma blocky_imp_unary: assumes "k \ length vs" and "blocky \ vs k" shows "unary \ vs = k" using assms satisfies_Psi Psi_unary by simp text \ The family $\Upsilon$ of CNF formulas also takes as parameter a list of variable indices. \ definition Upsilon :: "nat list \ formula" ("\") where "\ vs \ map (\s. [Pos s]) (take 2 vs) @ map (\s. [Neg s]) (drop 3 vs)" text \ For $|\mathit{vs}| > 2$, an assignment satisfies $\Upsilon(\mathit{vs})$ iff.\ it satisfies $\Psi(\mathit{vs}, 2)$ or $\Psi(\mathit{vs}, 3)$. \ lemma Psi_2_imp_Upsilon: fixes \ :: assignment assumes "\ \ \ vs 2" and "length vs > 2" shows "\ \ \ vs" proof - have "\ \ map (\s. [Pos s]) (take 2 vs)" using assms Psi_def satisfies_def by simp moreover have "\ \ map (\s. [Neg s]) (drop 3 vs)" using assms Psi_def satisfies_def - by (smt (z3) Cons_nth_drop_Suc One_nat_def Suc_1 Un_iff insert_iff list.set(2) list.simps(9) + by (smt (verit) Cons_nth_drop_Suc One_nat_def Suc_1 Un_iff insert_iff list.set(2) list.simps(9) numeral_3_eq_3 set_append) ultimately show ?thesis using Upsilon_def satisfies_def by auto qed lemma Psi_3_imp_Upsilon: assumes "\ \ \ vs 3" and "length vs > 2" shows "\ \ \ vs" proof - have "\ \ map (\s. [Pos s]) (take 2 vs)" using assms Psi_def satisfies_def by (metis eval_nat_numeral(3) map_append satisfies_append(1) take_Suc_conv_app_nth) moreover have "\ \ map (\s. [Neg s]) (drop 3 vs)" using assms Psi_def satisfies_def by simp ultimately show ?thesis using Upsilon_def satisfies_def by auto qed lemma Upsilon_imp_Psi_2_or_3: assumes "\ \ \ vs" and "length vs > 2" shows "\ \ \ vs 2 \ \ \ \ vs 3" proof - have "\ \ map (\s. [Pos s]) (take 2 vs)" using satisfies_def assms Upsilon_def by simp then have "satisfies_clause \ [Pos v]" if "v \ set (take 2 vs)" for v using satisfies_def that by simp then have "satisfies_literal \ (Pos v)" if "v \ set (take 2 vs)" for v using satisfies_clause_def that by simp then have 1: "\ v" if "v \ set (take 2 vs)" for v using that by simp then have 2: "satisfies_clause \ [Pos v]" if "v \ set (take 2 vs)" for v using that satisfies_clause_def by simp have "\ \ map (\s. [Neg s]) (drop 3 vs)" using satisfies_def assms Upsilon_def by simp then have "satisfies_clause \ [Neg v]" if "v \ set (drop 3 vs)" for v using satisfies_def that by simp then have "satisfies_literal \ (Neg v)" if "v \ set (drop 3 vs)" for v using satisfies_clause_def that by simp then have 3: "\ \ v" if "v \ set (drop 3 vs)" for v using that by simp then have 4: "satisfies_clause \ [Neg v]" if "v \ set (drop 3 vs)" for v using that satisfies_clause_def by simp show ?thesis proof (cases "\ (vs ! 2)") case True then have "\ v" if "v \ set (take 3 vs)" for v using that 1 assms(2) by (metis (no_types, lifting) in_set_conv_nth le_simps(3) length_take less_imp_le_nat linorder_neqE_nat min_absorb2 not_less_eq nth_take numeral_One numeral_plus_numeral plus_1_eq_Suc semiring_norm(3)) then have "satisfies_clause \ [Pos v]" if "v \ set (take 3 vs)" for v using that satisfies_clause_def by simp then have "\ \ \ vs 3" using 4 Psi_def satisfies_def by auto then show ?thesis by simp next case False then have "\ \ v" if "v \ set (drop 2 vs)" for v using that 3 assms(2) by (metis Cons_nth_drop_Suc numeral_plus_numeral numerals(1) plus_1_eq_Suc semiring_norm(3) set_ConsD) then have "satisfies_clause \ [Neg v]" if "v \ set (drop 2 vs)" for v using that satisfies_clause_def by simp then have "\ \ \ vs 2" using 2 Psi_def satisfies_def by auto then show ?thesis by simp qed qed lemma Upsilon_unary: assumes "\ \ \ vs" and "length vs > 2" shows "unary \ vs = 2 \ unary \ vs = 3" using assms Upsilon_imp_Psi_2_or_3 Psi_unary by fastforce section \The functions $\inputpos$ and $\prev$\ text \ Sequences of the symbol~\textbf{0}: \ definition zeros :: "nat \ symbol list" where "zeros n \ string_to_symbols (replicate n \)" lemma length_zeros [simp]: "length (zeros n) = n" using zeros_def by simp lemma bit_symbols_zeros: "bit_symbols (zeros n)" using zeros_def by simp lemma zeros: "zeros n = replicate n \" using zeros_def by simp text \ The assumptions in the following locale are the conditions that according to lemma~@{text NP_imp_oblivious_2tape} hold for all $\NP$ languages. The construction of $\Phi$ will take place inside this locale, which in later chapters will be extended to contain the Turing machine outputting $\Phi$ and the correctness proof for this Turing machine. \ locale reduction_sat = fixes L :: language fixes M :: machine and G :: nat and T p :: "nat \ nat" assumes T: "big_oh_poly T" assumes p: "polynomial p" assumes tm_M: "turing_machine 2 G M" and oblivious_M: "oblivious M" and T_halt: "\y. bit_symbols y \ fst (execute M (start_config 2 y) (T (length y))) = length M" and cert: "\x. x \ L \ (\u. length u = p (length x) \ execute M (start_config 2 \x; u\) (T (length \x; u\)) <.> 1 = \)" begin text \ The value $H$ is an upper bound for the number of states of $M$ and the alphabet size of $M$. \ definition H :: nat where "H \ max G (length M)" lemma H_ge_G: "H \ G" using H_def by simp lemma H_gr_2: "H > 2" using H_def tm_M turing_machine_def by auto lemma H_ge_3: "H \ 3" using H_gr_2 by simp lemma H_ge_length_M: "H \ length M" using H_def by simp text \ The number of symbols used for encoding one snapshot is $Z = 3H$: \ definition Z :: nat where "Z \ 3 * H" text \ The configuration after running $M$ on input $y$ for $t$ steps: \ abbreviation exc :: "symbol list \ nat \ config" where "exc y t \ execute M (start_config 2 y) t" text \ The function $T$ is just some polynomial upper bound for the running time. The next function, $TT$, is the actual running time. Since $M$ is oblivious, its running time depends only on the length of the input. The argument @{term "zeros n"} is thus merely a placeholder for an arbitrary symbol sequence of length $n$. \ definition TT :: "nat \ nat" where "TT n \ LEAST t. fst (exc (zeros n) t) = length M" lemma TT: "fst (exc (zeros n) (TT n)) = length M" proof - let ?P = "\t. fst (exc (zeros n) t) = length M" have "?P (T n)" using T_halt bit_symbols_zeros length_zeros by metis then show ?thesis using wellorder_Least_lemma[of ?P] TT_def by simp qed lemma TT_le: "TT n \ T n" using wellorder_Least_lemma length_zeros TT_def T_halt[OF bit_symbols_zeros[of n]] by fastforce lemma less_TT: "t < TT n \ fst (exc (zeros n) t) < length M" proof - assume "t < TT n" then have "fst (exc (zeros n) t) \ length M" using TT_def not_less_Least by auto moreover have "fst (exc (zeros n) t) \ length M" for t using tm_M start_config_def turing_machine_execute_states by auto ultimately show "fst (exc (zeros n) t) < length M" using less_le by blast qed lemma oblivious_halt_state: assumes "bit_symbols zs" shows "fst (exc zs t) < length M \ fst (exc (zeros (length zs)) t) < length M" proof - obtain e where e: "\zs. bit_symbols zs \ (\tps. trace M (start_config 2 zs) (e (length zs)) (length M, tps))" using oblivious_M oblivious_def by auto let ?es = "e (length zs)" have "\ii The position of the input tape head of $M$ depends only on the length $n$ of the input and the step $t$, at least as long as the input is over the alphabet $\{\mathbf{0}, \mathbf{1}\}$. \ definition inputpos :: "nat \ nat \ nat" where "inputpos n t \ exc (zeros n) t <#> 0" lemma inputpos_oblivious: assumes "bit_symbols zs" shows "exc zs t <#> 0 = inputpos (length zs) t" proof - obtain e where e: "(\zs. bit_symbols zs \ (\tps. trace M (start_config 2 zs) (e (length zs)) (length M, tps)))" using oblivious_M oblivious_def by auto let ?es = "e (length zs)" obtain tps where t1: "trace M (start_config 2 zs) ?es (length M, tps)" using e assms by auto let ?zs = "(replicate (length zs) 2)" have "proper_symbols ?zs" by simp moreover have "length ?zs = length zs" by simp ultimately obtain tps0 where t0: "trace M (start_config 2 ?zs) ?es (length M, tps0)" using e by fastforce have le: "exc zs t <#> 0 = inputpos (length zs) t" if "t \ length ?es" for t proof (cases "t = 0") case True then show ?thesis by (simp add: start_config_def inputpos_def) next case False then obtain i where i: "Suc i = t" using gr0_implies_Suc by auto then have "exc zs (Suc i) <#> 0 = fst (?es ! i)" using t1 False that Suc_le_lessD trace_def by auto moreover have "exc ?zs (Suc i) <#> 0 = fst (?es ! i)" using t0 False i that Suc_le_lessD trace_def by auto ultimately show ?thesis using i inputpos_def zeros by simp qed moreover have "exc zs t <#> 0 = inputpos (length zs) t" if "t > length ?es" proof - have "exc ?zs (length ?es) = (length M, tps0)" using t0 trace_def by simp then have *: "exc ?zs t = exc ?zs (length ?es)" using that by (metis execute_after_halting_ge fst_eqD less_or_eq_imp_le) have "exc zs (length ?es) = (length M, tps)" using t1 trace_def by simp then have "exc zs t = exc zs (length ?es)" using that by (metis execute_after_halting_ge fst_eqD less_or_eq_imp_le) then show ?thesis using * le[of "length ?es"] by (simp add: inputpos_def zeros) qed ultimately show ?thesis by fastforce qed text \ The position of the tape head on the output tape of $M$ also depends only on the length $n$ of the input and the step $t$. \ lemma oblivious_headpos_1: assumes "bit_symbols zs" shows "exc zs t <#> 1 = exc (zeros (length zs)) t <#> 1" proof - obtain e where e: "(\zs. bit_symbols zs \ (\tps. trace M (start_config 2 zs) (e (length zs)) (length M, tps)))" using oblivious_M oblivious_def by auto let ?es = "e (length zs)" obtain tps where t1: "trace M (start_config 2 zs) ?es (length M, tps)" using e assms by auto let ?zs = "(replicate (length zs) 2)" have "proper_symbols ?zs" by simp moreover have "length ?zs = length zs" by simp ultimately obtain tps0 where t0: "trace M (start_config 2 ?zs) ?es (length M, tps0)" using e by fastforce have le: "exc zs t <#> 1 = exc (zeros (length zs)) t <#> 1" if "t \ length ?es" for t proof (cases "t = 0") case True then show ?thesis by (simp add: start_config_def inputpos_def) next case False then obtain i where i: "Suc i = t" using gr0_implies_Suc by auto then have "exc zs (Suc i) <#> 1 = snd (?es ! i)" using t1 False that Suc_le_lessD trace_def by auto moreover have "exc ?zs (Suc i) <#> 1 = snd (?es ! i)" using t0 False i that Suc_le_lessD trace_def by auto ultimately show ?thesis using i inputpos_def zeros by simp qed moreover have "exc zs t <#> 1 = exc (zeros (length zs)) t <#> 1" if "t > length ?es" proof - have "exc ?zs (length ?es) = (length M, tps0)" using t0 trace_def by simp then have *: "exc ?zs t = exc ?zs (length ?es)" using that by (metis execute_after_halting_ge fst_eqD less_or_eq_imp_le) have "exc zs (length ?es) = (length M, tps)" using t1 trace_def by simp then have "exc zs t = exc zs (length ?es)" using that by (metis execute_after_halting_ge fst_eqD less_or_eq_imp_le) then show ?thesis using * le[of "length ?es"] by (simp add: inputpos_def zeros) qed ultimately show ?thesis using le_less_linear by blast qed text \ The value $\prev(t)$ is the most recent step in which $M$'s output tape head was in the same position as in step $t$. If no such step exists, $\prev(t)$ is set to $t$. Again due to $M$ being oblivious, $\prev$ depends only on the length $n$ of the input (and on $t$, of course). \ definition prev :: "nat \ nat \ nat" where "prev n t \ if \t' 1 = exc (zeros n) t <#> 1 then GREATEST t'. t' < t \ exc (zeros n) t' <#> 1 = exc (zeros n) t <#> 1 else t" lemma oblivious_prev: assumes "bit_symbols zs" shows "prev (length zs) t = (if \t' 1 = exc zs t <#> 1 then GREATEST t'. t' < t \ exc zs t' <#> 1 = exc zs t <#> 1 else t)" using prev_def assms oblivious_headpos_1 by auto lemma prev_less: assumes "\t' 1 = exc (zeros n) t <#> 1" shows "prev n t < t \ exc (zeros n) (prev n t) <#> 1 = exc (zeros n) t <#> 1" proof - let ?P = "\t'. t' < t \ exc (zeros n) t' <#> 1 = exc (zeros n) t <#> 1" have "prev n t = (GREATEST t'. t' < t \ exc (zeros n) t' <#> 1 = exc (zeros n) t <#> 1)" using assms prev_def by simp moreover have "\y. ?P y \ y \ t" by simp ultimately show ?thesis using GreatestI_ex_nat[OF assms, of t] by simp qed corollary prev_less': assumes "bit_symbols zs" assumes "\t' 1 = exc zs t <#> 1" shows "prev (length zs) t < t \ exc zs (prev (length zs) t) <#> 1 = exc zs t <#> 1" using prev_less oblivious_headpos_1 assms by simp lemma prev_greatest: assumes "t' < t" and "exc (zeros n) t' <#> 1 = exc (zeros n) t <#> 1" shows "t' \ prev n t" proof - let ?P = "\t'. t' < t \ exc (zeros n) t' <#> 1 = exc (zeros n) t <#> 1" have "prev n t = (GREATEST t'. t' < t \ exc (zeros n) t' <#> 1 = exc (zeros n) t <#> 1)" using assms prev_def by auto moreover have "?P t'" using assms by simp moreover have "\y. ?P y \ y \ t" by simp ultimately show ?thesis using Greatest_le_nat[of ?P t' t] by simp qed corollary prev_greatest': assumes "bit_symbols zs" assumes "t' < t" and "exc zs t' <#> 1 = exc zs t <#> 1" shows "t' \ prev (length zs) t" using prev_greatest oblivious_headpos_1 assms by simp lemma prev_eq: "prev n t = t \ \ (\t' 1 = exc (zeros n) t <#> 1)" using prev_def nat_less_le prev_less by simp lemma prev_le: "prev n t \ t" using prev_eq prev_less by (metis less_or_eq_imp_le) corollary prev_eq': assumes "bit_symbols zs" shows "prev (length zs) t = t \ \ (\t' 1 = exc zs t <#> 1)" using prev_eq oblivious_headpos_1 assms by simp lemma prev_between: assumes "prev n t < t'" and "t' < t" shows "exc (zeros n) t' <#> 1 \ exc (zeros n) (prev n t) <#> 1" using assms by (metis (no_types, lifting) leD prev_eq prev_greatest prev_less) lemma prev_write_read: assumes "bit_symbols zs" and "n = length zs" and "prev n t < t" and "cfg = exc zs (prev n t)" and "t \ TT n" shows "exc zs t <.> 1 = (M ! (fst cfg)) [cfg <.> 0, cfg <.> 1] [.] 1" proof - let ?cfg = "exc zs (Suc (prev n t))" let ?sas = "(M ! (fst cfg)) [cfg <.> 0, cfg <.> 1]" let ?i = "cfg <#> 1" have 1: "fst cfg < length M" using assms less_TT' by simp have 2: "||cfg|| = 2" using assms execute_num_tapes start_config_length tm_M by auto then have 3: "read (snd cfg) = [cfg <.> 0, cfg <.> 1]" using read_def - by (smt (z3) Cons_eq_map_conv Suc_1 length_0_conv length_Suc_conv list.simps(8) + by (smt (verit) Cons_eq_map_conv Suc_1 length_0_conv length_Suc_conv list.simps(8) nth_Cons_0 nth_Cons_Suc numeral_1_eq_Suc_0 numeral_One) have *: "(?cfg <:> 1) ?i = (M ! (fst cfg)) [cfg <.> 0, cfg <.> 1] [.] 1" proof - have "?cfg 1 = exe M cfg 1" by (simp add: assms) also have "... = sem (M ! (fst cfg)) cfg 1" using 1 exe_lt_length by simp also have "... = act (snd ((M ! (fst cfg)) (read (snd cfg))) ! 1) (snd cfg ! 1)" using sem_snd_tm tm_M 1 2 by (metis Suc_1 lessI prod.collapse) also have "... = act (?sas [!] 1) (cfg 1)" using 3 by simp finally have *: "?cfg 1 = act (?sas [!] 1) (cfg 1)" . have "(?cfg <:> 1) ?i = fst (?cfg 1) ?i" by simp also have ***: "... = ((fst (cfg 1))(?i := (?sas [.] 1))) ?i" using * act by simp also have "... = ?sas [.] 1" by simp finally show "(?cfg <:> 1) ?i = ?sas [.] 1" using *** by simp qed have "((exc zs t') <:> 1) ?i = (M ! (fst cfg)) [cfg <.> 0, cfg <.> 1] [.] 1" if "Suc (prev n t) \ t'" and "t' \ t" for t' using that proof (induction t' rule: nat_induct_at_least) case base then show ?case using * by simp next case (Suc m) let ?cfg_m = "exc zs m" from Suc have between: "prev n t < m" "m < t" by simp_all then have *: "?cfg_m <#> 1 \ ?i" using prev_between assms oblivious_headpos_1 by simp have "m < TT n" using Suc assms by simp then have 1: "fst ?cfg_m < length M" using assms less_TT' by simp have 2: "||?cfg_m|| = 2" using execute_num_tapes start_config_length tm_M by auto have "exc zs (Suc m) 1 = exe M ?cfg_m 1" by simp also have "... = sem (M ! (fst ?cfg_m)) ?cfg_m 1" using 1 exe_lt_length by simp also have "... = act (snd ((M ! (fst ?cfg_m)) (read (snd ?cfg_m))) ! 1) (snd ?cfg_m ! 1)" using sem_snd_tm tm_M 1 2 by (metis Suc_1 lessI prod.collapse) finally have "exc zs (Suc m) 1 = act (snd ((M ! (fst ?cfg_m)) (read (snd ?cfg_m))) ! 1) (snd ?cfg_m ! 1)" . then have "(exc zs (Suc m) <:> 1) ?i = fst (snd ?cfg_m ! 1) ?i" using * act_changes_at_most_pos' by simp then show ?case using Suc by simp qed then have "((exc zs t) <:> 1) ?i = (M ! (fst cfg)) [cfg <.> 0, cfg <.> 1] [.] 1" using Suc_leI assms by simp moreover have "?i = exc zs t <#> 1" - using assms(1,2,4) oblivious_headpos_1 prev_eq prev_less by (smt (z3)) + using assms(1,2,4) oblivious_headpos_1 prev_eq prev_less by (smt (verit)) ultimately show ?thesis by simp qed lemma prev_no_write: assumes "bit_symbols zs" and "n = length zs" and "prev n t = t" and "t \ TT n" and "t > 0" shows "exc zs t <.> 1 = \" proof - let ?i = "exc zs t <#> 1" have 1: "\ (\t' 1 = ?i)" using prev_eq' assms(1,2,3) by simp have 2: "?i > 0" proof (rule ccontr) assume "\ 0 < ?i" then have eq0: "?i = 0" by simp moreover have "exc zs 0 <#> 1 = 0" by (simp add: start_config_pos) ultimately show False using 1 assms(5) by auto qed have 3: "(exc zs (Suc t') <:> 1) i = (exc zs t' <:> 1) i" if "i \ exc zs t' <#> 1" for i t' proof (cases "fst (exc zs t') < length M") case True let ?cfg = "exc zs t'" have len2: "||?cfg|| = 2" using execute_num_tapes start_config_length tm_M by auto have "exc zs (Suc t') 1 = exe M ?cfg 1" by simp also have "... = sem (M ! (fst ?cfg)) ?cfg 1" using True exe_lt_length by simp also have "... = act (snd ((M ! (fst ?cfg)) (read (snd ?cfg))) ! 1) (snd ?cfg ! 1)" using sem_snd_tm tm_M True len2 by (metis Suc_1 lessI prod.collapse) finally have "exc zs (Suc t') 1 = act (snd ((M ! (fst ?cfg)) (read (snd ?cfg))) ! 1) (snd ?cfg ! 1)" . then have "(exc zs (Suc t') <:> 1) i = fst (snd ?cfg ! 1) i" using that act_changes_at_most_pos' by simp then show ?thesis by simp next case False then show ?thesis using that by (simp add: exe_def) qed have "(exc zs t' <:> 1) ?i = (exc zs 0 <:> 1) ?i" if "t' \ t" for t' using that proof (induction t') case 0 then show ?case by simp next case (Suc t') then show ?case by (metis 1 3 Suc_leD Suc_le_lessD) qed then have "exc zs t <.> 1 = (exc zs 0 <:> 1) ?i" by simp then show ?thesis using 2 One_nat_def execute.simps(1) start_config1 less_2_cases_iff less_one by presburger qed text \ The intervals $\gamma_i$ and $w_0, \dots, w_9$ do not depend on $x$, and so can be defined here already. \ definition gamma :: "nat \ nat list" ("\") where "\ i \ [i * H.. i) = H" using gamma_def by simp abbreviation "w\<^sub>0 \ [0..1 \ [H..<2*H]" abbreviation "w\<^sub>2 \ [2*H..3 \ [Z..4 \ [Z+H..5 \ [Z+2*H..<2*Z]" abbreviation "w\<^sub>6 \ [2*Z..<2*Z+H]" abbreviation "w\<^sub>7 \ [2*Z+H..<2*Z+2*H]" abbreviation "w\<^sub>8 \ [2*Z+2*H..<3*Z]" abbreviation "w\<^sub>9 \ [3*Z..<3*Z+H]" lemma unary_upt_eq: fixes \\<^sub>1 \\<^sub>2 :: assignment and lower upper k :: nat assumes "\i\<^sub>1 i = \\<^sub>2 i" and "upper \ k" shows "unary \\<^sub>1 [lower..\<^sub>2 [lower..\<^sub>1 [lower..\<^sub>2 [lower..\<^sub>1 ?vs = filter \\<^sub>2 ?vs" using assms by (metis atLeastLessThan_iff filter_cong less_le_trans set_upt) then show ?thesis using numtrue_def by simp qed moreover have "blocky \\<^sub>1 [lower..\<^sub>2 [lower.. For the case @{term "prev m t < t"}, we have the following predicate on assignments, which corresponds to ($\mathrm{F}_1$), ($\mathrm{F}_2$), ($\mathrm{F}_3$) from the introduction: \ definition F :: "assignment \ bool" where "F \ \ unary \ w\<^sub>6 = unary \ w\<^sub>9 \ unary \ w\<^sub>7 = (M ! (unary \ w\<^sub>5)) [unary \ w\<^sub>3, unary \ w\<^sub>4] [.] 1 \ unary \ w\<^sub>8 = fst ((M ! (unary \ w\<^sub>2)) [unary \ w\<^sub>0, unary \ w\<^sub>1])" lemma depon_F: "depon (3 * Z + H) F" using depon_def F_def Z_def unary_upt_eq by simp text \ There is a CNF formula $\psi$ that contains the first $3Z + H$ variables and is satisfied by exactly the assignments specified by $F$. \ definition psi :: formula ("\") where "\ \ SOME \. fsize \ \ (3 * Z + H) * 2 ^ (3 * Z + H) \ length \ \ 2 ^ (3 * Z + H) \ variables \ \ {..<3 * Z + H} \ (\\. F \ = \ \ \)" lemma psi: "fsize \ \ (3 * Z + H) * 2 ^ (3*Z + H) \ length \ \ 2 ^ (3 * Z + H) \ variables \ \ {..<3 * Z + H} \ (\\. F \ = \ \ \)" using psi_def someI_ex[OF depon_ex_formula[OF depon_F]] by metis lemma satisfies_psi: assumes "length \ = 3 * Z + H" shows "\ \ relabel \ \ = remap \ \ \ \" using assms psi satisfies_sigma by simp lemma psi_F: "remap \ \ \ \ = F (remap \ \)" using psi by simp corollary satisfies_F: assumes "length \ = 3 * Z + H" shows "\ \ relabel \ \ = F (remap \ \)" using assms satisfies_psi psi_F by simp text \ For the case @{term "prev m t = t"}, the following predicate corresponds to ($\mathrm{F}_1'$), ($\mathrm{F}_2'$), ($\mathrm{F}_3'$) from the introduction: \ definition F' :: "assignment \ bool" where "F' \ \ unary \ w\<^sub>3 = unary \ w\<^sub>6 \ unary \ w\<^sub>4 = 0 \ unary \ w\<^sub>5 = fst ((M ! (unary \ w\<^sub>2)) [unary \ w\<^sub>0, unary \ w\<^sub>1])" lemma depon_F': "depon (2 * Z + H) F'" using depon_def F'_def Z_def unary_upt_eq by simp text \ The CNF formula $\psi'$ is analogous to $\psi$ from the previous case. \ definition psi' :: formula ("\''") where "\' \ SOME \. fsize \ \ (2*Z+H) * 2 ^ (2*Z+H) \ length \ \ 2 ^ (2*Z+H) \ variables \ \ {..<2*Z+H} \ (\\. F' \ = \ \ \)" lemma psi': "fsize \' \ (2*Z+H) * 2 ^ (2*Z+H) \ length \' \ 2 ^ (2*Z+H) \ variables \' \ {..<2*Z+H} \ (\\. F' \ = \ \ \')" using psi'_def someI_ex[OF depon_ex_formula[OF depon_F']] by metis lemma satisfies_psi': assumes "length \ = 2*Z+H" shows "\ \ relabel \ \' = remap \ \ \ \'" using assms psi' satisfies_sigma by simp lemma psi'_F': "remap \ \ \ \' = F' (remap \ \)" using psi' by simp corollary satisfies_F': assumes "length \ = 2*Z+H" shows "\ \ relabel \ \' = F' (remap \ \)" using assms satisfies_psi' psi'_F' by simp end (* locale reduction_sat *) section \Snapshots\ text \ The snapshots and much of the rest of the construction of $\Phi$ depend on the string $x$. We encapsulate this in a sublocale of @{locale reduction_sat}. \ locale reduction_sat_x = reduction_sat + fixes x :: string begin abbreviation n :: nat where "n \ length x" text \ Turing machines consume the string $x$ as a sequence $xs$ of symbols: \ abbreviation xs :: "symbol list" where "xs \ string_to_symbols x" lemma bs_xs: "bit_symbols xs" by simp text \ For the verifier Turing machine $M$ we are only concerned with inputs of the form $\langle x, u\rangle$ for a string $u$ of length $p(n)$. The pair $\langle x, u\rangle$ has the length $m = 2n + 2p(n) + 2$. \ definition m :: nat where "m \ 2 * n + 2 * p n + 2" text \ On input $\langle x, u\rangle$ the Turing machine $M$ halts after $T' = TT(m)$ steps. \ definition T' :: nat where "T' \ TT m" text \ The positions of both of $M$'s tape heads are bounded by $T'$. \ lemma inputpos_less: "inputpos m t \ T'" proof - define u :: string where "u = replicate (p n) False" let ?i = "inputpos m t" have y: "bit_symbols \x; u\" by simp have len: "length \x; u\ = m" using u_def m_def length_pair by simp then have "exc \x; u\ t <#> 0 \ T'" using TT'[OF y] T'_def head_pos_le_halting_time[OF tm_M, of "\x; u\" T' 0] by simp then show ?thesis using inputpos_oblivious[OF y] len by simp qed lemma headpos_1_less: "exc (zeros m) t <#> 1 \ T'" proof - define u :: string where "u = replicate (p n) False" let ?i = "inputpos m t" have y: "bit_symbols \x; u\" by simp have len: "length \x; u\ = m" using u_def m_def length_pair by simp then have "exc \x; u\ t <#> 1 \ T'" using TT'[OF y] T'_def head_pos_le_halting_time[OF tm_M, of "\x; u\" "T'" 1] by simp then show ?thesis using oblivious_headpos_1[OF y] len by simp qed text \ The formula $\Phi$ must contain a condition for every symbol that $M$ is reading from the input tape. While $T'$ is an upper bound for the input tape head position of $M$, it might be that $T'$ is less than the length of the input $\langle x, u\rangle$. So the portion of the input read by $M$ might be a prefix of the input or it might be the input followed by some blanks afterwards. This would make for an awkward case distinction. We do not have to be very precise here and can afford to bound the portion of the input tape read by $M$ by the number $m' = 2n + 2p(n) + 3 + T'$, which is the length of the start symbol followed by the input $\langle x, u\rangle$ followed by $T'$ blanks. This symbol sequence was called $y(u)$ in the introduction. Here we will call it @{term "ysymbols u"}. \ definition m' :: nat where "m' \ 2 * n + 2 * p n + 3 + T'" definition ysymbols :: "string \ symbol list" where "ysymbols u \ 1 # \x; u\ @ replicate T' 0" lemma length_ysymbols: "length u = p n \ length (ysymbols u) = m'" using ysymbols_def m'_def m_def length_pair by simp lemma ysymbols_init: assumes "i < length (ysymbols u)" shows "ysymbols u ! i = (start_config 2 \x; u\ <:> 0) i" proof - let ?y = "\x; u\" have init: "start_config 2 ?y <:> 0 = (\i. if i = 0 then 1 else if i \ length ?y then ?y ! (i - 1) else 0)" using start_config_def by auto have "i < length ?y + 1 + T'" using assms ysymbols_def by simp then consider "i = 0" | "i > 0 \ i \ length ?y" | "i > length ?y \ i < length ?y + 1 + T'" by linarith then show "ysymbols u ! i = (start_config 2 ?y <:> 0) i" proof (cases) case 1 then show ?thesis using ysymbols_def init by simp next case 2 then have "ysymbols u ! i = \x; u\ ! (i - 1)" using ysymbols_def - by (smt (z3) One_nat_def Suc_less_eq Suc_pred le_imp_less_Suc neq0_conv nth_Cons' nth_append) + by (smt (verit, del_insts) Nat.add_diff_assoc diff_is_0_eq gr_zeroI le_add_diff_inverse le_add_diff_inverse2 less_numeral_extra(1) nat_le_linear nth_Cons_pos nth_append zero_less_diff) then show ?thesis using init 2 by simp next case 3 then have "(start_config 2 ?y <:> 0) i = 0" using init by simp moreover have "ysymbols u ! i = 0" unfolding ysymbols_def using 3 nth_append[of "1#\x; u\" "replicate T' 0" i] by auto ultimately show ?thesis by simp qed qed lemma ysymbols_at_0: "ysymbols u ! 0 = 1" using ysymbols_def by simp lemma ysymbols_first_at: assumes "j < length x" shows "ysymbols u ! (2*j+1) = 2" and "ysymbols u ! (2*j+2) = (if x ! j then 3 else 2)" proof - have *: "ysymbols u = (1 # \x; u\) @ replicate T' 0" using ysymbols_def by simp let ?i = "2 * j + 1" have len: "2*j < length \x, u\" using assms length_string_pair by simp have "?i < length (1 # \x; u\)" using assms length_pair by simp then have "ysymbols u ! ?i = (1 # \x; u\) ! ?i" using * nth_append by metis also have "... = \x; u\ ! (2*j)" by simp also have "... = 2" using string_pair_first_nth assms len by simp finally show "ysymbols u ! ?i = 2" . let ?i = "2 * j + 2" have len2: "?i < length (1 # \x; u\)" using assms length_pair by simp then have "ysymbols u ! ?i = (1 # \x; u\) ! ?i" using * nth_append by metis also have "... = \x; u\ ! (2*j+1)" by simp also have "... = (if x ! j then 3 else 2)" using string_pair_first_nth(2) assms len2 by simp finally show "ysymbols u ! ?i = (if x ! j then 3 else 2)" . qed lemma ysymbols_at_2n1: "ysymbols u ! (2*n+1) = 3" proof - let ?i = "2 * n + 1" have "ysymbols u ! ?i = \x; u\ ! (2*n)" using ysymbols_def by (metis (no_types, lifting) add.commute add_2_eq_Suc' le_add2 le_imp_less_Suc length_pair less_SucI nth_Cons_Suc nth_append plus_1_eq_Suc) also have "... = (if \x, u\ ! (2*n) then 3 else 2)" using length_pair by simp also have "... = 3" using string_pair_sep_nth by simp finally show ?thesis . qed lemma ysymbols_at_2n2: "ysymbols u ! (2*n+2) = 3" proof - let ?i = "2 * n + 2" have "ysymbols u ! ?i = \x; u\ ! (2*n+1)" - using ysymbols_def - by (smt (z3) One_nat_def add.commute add.right_neutral add_2_eq_Suc' add_Suc_right le_add2 - le_imp_less_Suc length_pair nth_Cons_Suc nth_append) + by (simp add: ysymbols_def) + (smt (verit, del_insts) add.right_neutral add_2_eq_Suc' length_greater_0_conv length_pair + lessI less_add_same_cancel1 less_trans_Suc list.size(3) mult_0_right mult_pos_pos nth_append + zero_less_numeral) also have "... = (if \x, u\ ! (2*n+1) then 3 else 2)" using length_pair by simp also have "... = 3" using string_pair_sep_nth by simp finally show ?thesis . qed lemma ysymbols_second_at: assumes "j < length u" shows "ysymbols u ! (2*n+2+2*j+1) = 2" and "ysymbols u ! (2*n+2+2*j+2) = (if u ! j then 3 else 2)" proof - have *: "ysymbols u = (1 # \x; u\) @ replicate T' 0" using ysymbols_def by simp let ?i = "2 * n + 2 + 2 * j + 1" have len: "2 * n + 2 + 2*j < length \x, u\" using assms length_string_pair by simp have "?i < length (1 # \x; u\)" using assms length_pair by simp then have "ysymbols u ! ?i = (1 # \x; u\) ! ?i" using * nth_append by metis also have "... = \x; u\ ! (2*n+2+2*j)" by simp also have "... = 2" using string_pair_second_nth(1) assms len by simp finally show "ysymbols u ! ?i = 2" . let ?i = "2*n+2+2 * j + 2" have len2: "?i < length (1 # \x; u\)" using assms length_pair by simp then have "ysymbols u ! ?i = (1 # \x; u\) ! ?i" using * nth_append by metis also have "... = \x; u\ ! (2*n+2+2*j+1)" by simp also have "... = (if u ! j then 3 else 2)" using string_pair_second_nth(2) assms len2 by simp finally show "ysymbols u ! ?i = (if u ! j then 3 else 2)" . qed lemma ysymbols_last: assumes "length u = p n" and "i > m" and "i < m + 1 + T'" shows "ysymbols u ! i = 0" using assms length_ysymbols m_def m'_def ysymbols_def nth_append[of "1#\x; u\" "replicate T' 0" i] by simp text \ The number of symbols used for unary encoding $m'$ symbols will be called $N$: \ definition N :: nat where "N \ H * m'" lemma N_eq: "N = H * (2 * n + 2 * p n + 3 + T')" using m'_def N_def by simp lemma m': "m' * H = N " using m'_def N_def by simp lemma inputpos_less': "inputpos m t < m'" using inputpos_less m_def m'_def by (metis Suc_1 add_less_mono1 le_neq_implies_less lessI linorder_neqE_nat not_add_less2 numeral_plus_numeral semiring_norm(3) trans_less_add2) lemma T'_less: "T' < N" proof - have "T' < 2 * n + 2 * p n + 3 + T'" by simp also have "... < H * (2 * n + 2 * p n + 3 + T')" using H_gr_2 by simp also have "... = N" using N_eq by simp finally show ?thesis . qed text \The three components of a snapshot:\ definition z0 :: "string \ nat \ symbol" where "z0 u t \ exc \x; u\ t <.> 0" definition z1 :: "string \ nat \ symbol" where "z1 u t \ exc \x; u\ t <.> 1" definition z2 :: "string \ nat \ state" where "z2 u t \ fst (exc \x; u\ t)" lemma z0_le: "z0 u t \ H" using z0_def H_ge_G tape_alphabet[OF tm_M, where ?j=0 and ?zs="\x; u\"] symbols_lt_pair[of x u] tm_M turing_machine_def by (metis (no_types, lifting) dual_order.strict_trans1 less_or_eq_imp_le zero_less_numeral) lemma z1_le: "z1 u t \ H" using z1_def H_ge_G tape_alphabet[OF tm_M, where ?j=1 and ?zs="\x; u\"] symbols_lt_pair[of x u] tm_M turing_machine_def by (metis (no_types, lifting) dual_order.strict_trans1 less_or_eq_imp_le one_less_numeral_iff semiring_norm(76)) lemma z2_le: "z2 u t \ H" proof - have "z2 u t \ length M" using z2_def turing_machine_execute_states[OF tm_M] start_config_def by simp then show ?thesis using H_ge_length_M by simp qed text \ The next lemma corresponds to (Z1) from the second equivalence mentioned in the introduction. It expresses the first element of a snapshot in terms of $y(u)$ and $\inputpos$. \ lemma z0: assumes "length u = p n" shows "z0 u t = ysymbols u ! (inputpos m t)" proof - let ?i = "inputpos m t" let ?y = "\x; u\" have "bit_symbols ?y" by simp have len: "length ?y = m" using assms m_def length_pair by simp have "?i < length (ysymbols u)" using inputpos_less' assms length_ysymbols by simp then have *: "ysymbols u ! ?i = (start_config 2 ?y <:> 0) ?i" using ysymbols_init by simp have "z0 u t = exc ?y t <.> 0" using z0_def by simp also have "... = (start_config 2 ?y <:> 0) (exc ?y t <#> 0)" using tm_M input_tape_constant start_config_length by simp also have "... = (start_config 2 ?y <:> 0) ?i" using inputpos_oblivious[OF `bit_symbols ?y`] len by simp also have "... = ysymbols u ! ?i" using * by simp finally show ?thesis . qed text \ The next lemma corresponds to (Z2) from the second equivalence mentioned in the introduction. It shows how, in the case $\prev(t) < t$, the second component of a snapshot can be expressed recursively using snapshots for earlier steps. \ lemma z1: assumes "length u = p n" and "prev m t < t" and "t \ T'" shows "z1 u t = (M ! (z2 u (prev m t))) [z0 u (prev m t), z1 u (prev m t)] [.] 1" proof - let ?y = "\x; u\" let ?cfg = "exc ?y (prev m t)" have "bit_symbols ?y" by simp moreover have len: "length ?y = m" using assms m_def length_pair by simp ultimately have "exc ?y t <.> 1 = (M ! (fst ?cfg)) [?cfg <.> 0, ?cfg <.> 1] [.] 1" using prev_write_read[of ?y m t ?cfg] assms(2,3) T'_def by fastforce then show ?thesis using z0_def z1_def z2_def by simp qed text \ The next lemma corresponds to (Z3) from the second equivalence mentioned in the introduction. It shows that in the case $\prev(t) = t$, the second component of a snapshot equals the blank symbol. \ lemma z1': assumes "length u = p n" and "prev m t = t" and "0 < t" and "t \ T'" shows "z1 u t = \" proof - let ?y = "\x; u\" let ?cfg = "exc ?y (prev m t)" have "bit_symbols ?y" by simp moreover have len: "length ?y = m" using assms m_def length_pair by simp ultimately have "exc ?y t <.> 1 = \" using prev_no_write[of ?y m t] assms T'_def by fastforce then show ?thesis using z0_def z1_def z2_def by simp qed text \ The next lemma corresponds to (Z4) from the second equivalence mentioned in the introduction. It shows how the third component of a snapshot can be expressed recursively using snapshots for earlier steps. \ lemma z2: assumes "length u = p n" and "t < T'" shows "z2 u (Suc t) = fst ((M ! (z2 u t)) [z0 u t, z1 u t])" proof - let ?y = "\x; u\" have "bit_symbols ?y" by simp moreover have len: "length ?y = m" using assms m_def length_pair by simp ultimately have run: "fst (exc ?y t) < length M" using less_TT' assms(2) T'_def by simp have "||exc ?y t|| = 2" using start_config_length execute_num_tapes tm_M by simp then have "snd (exc ?y t) = [exc ?y t 0, exc ?y t 1]" - by auto (smt (z3) Suc_length_conv length_0_conv nth_Cons_0 nth_Cons_Suc numeral_2_eq_2) + by auto (smt (verit) Suc_length_conv length_0_conv nth_Cons_0 nth_Cons_Suc numeral_2_eq_2) then have *: "read (snd (exc ?y t)) = [exc ?y t <.> 0, exc ?y t <.> 1]" using read_def by (metis (no_types, lifting) list.simps(8) list.simps(9)) have "z2 u (Suc t) = fst (exc ?y (Suc t))" using z2_def by simp also have "... = fst (exe M (exc ?y t))" by simp also have "... = fst (sem (M ! fst (exc ?y t)) (exc ?y t))" using exe_lt_length run by simp also have "... = fst (sem (M ! (z2 u t)) (exc ?y t))" using z2_def by simp also have "... = fst ((M ! (z2 u t)) (read (snd (exc ?y t))))" using sem_fst by simp also have "... = fst ((M ! (z2 u t)) [exc ?y t <.> 0, exc ?y t <.> 1])" using * by simp also have "... = fst ((M ! (z2 u t)) [z0 u t, z1 u t])" using z0_def z1_def by simp finally show ?thesis . qed corollary z2': assumes "length u = p n" and "t > 0" and "t \ T'" shows "z2 u t = fst ((M ! (z2 u (t - 1))) [z0 u (t - 1), z1 u (t - 1)])" using assms z2 by (metis Suc_diff_1 Suc_less_eq le_imp_less_Suc) text \ The intervals $\zeta_0$, $\zeta_1$, and $\zeta_2$ are long enough for a unary encoding of the three components of a snapshot: \ definition zeta0 :: "nat \ nat list" ("\\<^sub>0") where "\\<^sub>0 t \ [N + t * Z.. nat list" ("\\<^sub>1") where "\\<^sub>1 t \ [N + t * Z + H.. nat list" ("\\<^sub>2") where "\\<^sub>2 t \ [N + t * Z + 2 * H..\<^sub>0 t) = H" using zeta0_def by simp lemma length_zeta1 [simp]: "length (\\<^sub>1 t) = H" using zeta1_def by simp lemma length_zeta2 [simp]: "length (\\<^sub>2 t) = H" using zeta2_def Z_def by simp text \ The substitutions $\rho_t$, which have to be applied to $\psi$ to get the CNF formulas $\chi_t$ for the case $\prev(t) < t$: \ definition rho :: "nat \ nat list" ("\") where "\ t \ \\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" lemma length_rho: "length (\ t) = 3 * Z + H" using rho_def Z_def by simp text \ The substitutions $\rho_t'$, which have to be applied to $\psi'$ to get the CNF formulas $\chi_t$ for the case $\prev(t) = t$: \ definition rho' :: "nat \ nat list" ("\''") where "\' t \ \\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" lemma length_rho': "length (\' t) = 2 * Z + H" using rho'_def Z_def by simp text \ An auxiliary lemma for accessing the $n$-th element of a list sandwiched between two lists. It will be applied to $xs = \rho_t$ or $xs = \rho'_t$: \ lemma nth_append3: fixes xs ys zs ws :: "'a list" and n i :: nat assumes "xs = ys @ zs @ ws" and "i < length zs" and "n = length ys" shows "xs ! (n + i) = zs ! i" using assms by (simp add: nth_append) text \The formulas $\chi_t$ representing snapshots for $0 < t \leq T'$:\ definition chi :: "nat \ formula" ("\") where "\ t \ if prev m t < t then relabel (\ t) \ else relabel (\' t) \'" text \ The crucial feature of the formulas $\chi_t$ for $t > 0$ is that they are satisfied by exactly those assignments that represent in their bits $N$ to $N + Z\cdot(T' + 1)$ the $T' + 1$ snapshots of $M$ on input $\langle x, u\rangle$ when the relevant portion of the input tape is encoded in the first $N$ bits of the assignment. This works because the $\chi_t$ constrain the assignment to meet the recursive characterizations (Z1) --- (Z4) for the snapshots. The next two lemmas make this more precise. We first consider the case $\prev(t) < t$. The following lemma says $\alpha$ satisfies $\chi_t$ iff.\ $\alpha$ satisfies the properties (Z1), (Z2), and (Z4). \ lemma satisfies_chi_less: fixes \ :: assignment assumes "prev m t < t" shows "\ \ \ t \ unary \ (\\<^sub>0 t) = unary \ (\ (inputpos m t)) \ unary \ (\\<^sub>1 t) = (M ! (unary \ (\\<^sub>2 (prev m t)))) [unary \ (\\<^sub>0 (prev m t)), unary \ (\\<^sub>1 (prev m t))] [.] 1 \ unary \ (\\<^sub>2 t) = fst ((M ! (unary \ (\\<^sub>2 (t - 1)))) [unary \ (\\<^sub>0 (t - 1)), unary \ (\\<^sub>1 (t - 1))])" proof - let ?sigma = "\ t" have "\ \ \ t = \ \ relabel ?sigma psi" using assms chi_def by simp then have "\ \ \ t = F (remap ?sigma \)" (is "_ = F ?alpha") by (simp add: length_rho satisfies_F) then have *: "\ \ \ t = (unary ?alpha w\<^sub>6 = unary ?alpha w\<^sub>9 \ unary ?alpha w\<^sub>7 = (M ! (unary ?alpha w\<^sub>5)) [unary ?alpha w\<^sub>3, unary ?alpha w\<^sub>4] [.] 1 \ unary ?alpha w\<^sub>8 = fst ((M ! (unary ?alpha w\<^sub>2)) [unary ?alpha w\<^sub>0, unary ?alpha w\<^sub>1]))" using F_def by simp have w_less_len_rho: "\s\set w\<^sub>0. s < length (\ t)" "\s\set w\<^sub>1. s < length (\ t)" "\s\set w\<^sub>2. s < length (\ t)" "\s\set w\<^sub>3. s < length (\ t)" "\s\set w\<^sub>4. s < length (\ t)" "\s\set w\<^sub>5. s < length (\ t)" "\s\set w\<^sub>6. s < length (\ t)" "\s\set w\<^sub>7. s < length (\ t)" "\s\set w\<^sub>8. s < length (\ t)" "\s\set w\<^sub>9. s < length (\ t)" using length_rho Z_def by simp_all have **: "\ \ \ t = (unary \ (reseq ?sigma w\<^sub>6) = unary \ (reseq ?sigma w\<^sub>9) \ unary \ (reseq ?sigma w\<^sub>7) = (M ! (unary \ (reseq ?sigma w\<^sub>5))) [unary \ (reseq ?sigma w\<^sub>3), unary \ (reseq ?sigma w\<^sub>4)] [.] 1 \ unary \ (reseq ?sigma w\<^sub>8) = fst ((M ! (unary \ (reseq ?sigma w\<^sub>2))) [unary \ (reseq ?sigma w\<^sub>0), unary \ (reseq ?sigma w\<^sub>1)]))" using * w_less_len_rho unary_remap[where ?\="?sigma" and ?\=\] by presburger have 1: "reseq ?sigma w\<^sub>0 = \\<^sub>0 (t - 1)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta0_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = [] @ \\<^sub>0 (t - 1) @ (\\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t))" using rho_def by simp have "?sigma ! i = \\<^sub>0 (t - 1) ! i" using nth_append3[OF 1, of i 0] Z_def that by simp then show ?thesis using reseq_def that by simp qed qed have 2: "reseq ?sigma w\<^sub>1 = \\<^sub>1 (t - 1)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta1_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = \\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (H+i) = \\<^sub>1 (t - 1) ! i" using that zeta0_def zeta1_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 3: "reseq ?sigma w\<^sub>2 = \\<^sub>2 (t - 1)" (is "?l = ?r") proof (rule nth_equalityI) show len: "length ?l = length ?r" using zeta2_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1)) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (2*H+i) = \\<^sub>2 (t - 1) ! i" using len that zeta0_def zeta1_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 4: "reseq ?sigma w\<^sub>3 = \\<^sub>0 (prev m t)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta0_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1)) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (Z+i) = \\<^sub>0 (prev m t) ! i" using that Z_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 5: "reseq ?sigma w\<^sub>4 = \\<^sub>1 (prev m t)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta1_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t)) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (Z+H+i) = \\<^sub>1 (prev m t) ! i" using that Z_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 6: "reseq ?sigma w\<^sub>5 = \\<^sub>2 (prev m t)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta2_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t)) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (Z+2*H+i) = \\<^sub>2 (prev m t) ! i" using that zeta0_def zeta1_def zeta2_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 7: "reseq ?sigma w\<^sub>6 = \\<^sub>0 t" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta0_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t)) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (2*Z+i) = \\<^sub>0 t ! i" using that Z_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 8: "reseq ?sigma w\<^sub>7 = \\<^sub>1 t" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta1_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t) @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (2*Z+H+i) = \\<^sub>1 t ! i" using that Z_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 9: "reseq ?sigma w\<^sub>8 = \\<^sub>2 t" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta2_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t) @ \\<^sub>2 t @ \ (inputpos m t)" using rho_def by simp have "?sigma ! (2*Z+2*H+i) = \\<^sub>2 t ! i" using that zeta2_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 10: "reseq ?sigma w\<^sub>9 = \ (inputpos m t)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using gamma_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t) @ \ (inputpos m t) @ []" using rho_def by simp have "?sigma ! (3*Z+i) = \ (inputpos m t) ! i" using that Z_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed show ?thesis using ** 1 2 3 4 5 6 7 8 9 10 by simp qed text \ Next we consider the case $\prev(t) = t$. The following lemma says $\alpha$ satisfies $\chi_t$ iff.\ $\alpha$ satisfies the properties (Z1), (Z2), and (Z3). \ lemma satisfies_chi_eq: assumes "prev m t = t" and "t \ T'" shows "\ \ \ t \ unary \ (\\<^sub>0 t) = unary \ (\ (inputpos m t)) \ unary \ (\\<^sub>1 t) = 0 \ unary \ (\\<^sub>2 t) = fst ((M ! (unary \ (\\<^sub>2 (t - 1)))) [unary \ (\\<^sub>0 (t - 1)), unary \ (\\<^sub>1 (t - 1))])" proof - let ?sigma = "\' t" have "\ \ \ t = \ \ relabel ?sigma \'" using assms(1) chi_def by simp then have "\ \ \ t = F' (remap ?sigma \)" (is "_ = F' ?alpha") by (simp add: length_rho' satisfies_F') then have *: "\ \ \ t = (unary ?alpha w\<^sub>3 = unary ?alpha w\<^sub>6 \ unary ?alpha w\<^sub>4 = 0 \ unary ?alpha w\<^sub>5 = fst ((M ! (unary ?alpha w\<^sub>2)) [unary ?alpha w\<^sub>0, unary ?alpha w\<^sub>1]))" using F'_def by simp have w_less_len_rho': "\s\set w\<^sub>0. s < length (\' t)" "\s\set w\<^sub>1. s < length (\' t)" "\s\set w\<^sub>2. s < length (\' t)" "\s\set w\<^sub>3. s < length (\' t)" "\s\set w\<^sub>4. s < length (\' t)" "\s\set w\<^sub>5. s < length (\' t)" "\s\set w\<^sub>6. s < length (\' t)" using length_rho' Z_def by simp_all have **: "\ \ \ t = (unary \ (reseq ?sigma w\<^sub>3) = unary \ (reseq ?sigma w\<^sub>6) \ unary \ (reseq ?sigma w\<^sub>4) = 0 \ unary \ (reseq ?sigma w\<^sub>5) = fst ((M ! (unary \ (reseq ?sigma w\<^sub>2))) [unary \ (reseq ?sigma w\<^sub>0), unary \ (reseq ?sigma w\<^sub>1)]))" using * w_less_len_rho' unary_remap[where ?\="?sigma" and ?\=\] by presburger have 1: "reseq ?sigma w\<^sub>0 = \\<^sub>0 (t - 1)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta0_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = [] @ \\<^sub>0 (t - 1) @ (\\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t))" using rho'_def by simp have "?sigma ! i = \\<^sub>0 (t - 1) ! i" using nth_append3[OF 1, of i 0] Z_def that by simp then show ?thesis using reseq_def that by simp qed qed have 2: "reseq ?sigma w\<^sub>1 = \\<^sub>1 (t - 1)" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta1_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = \\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho'_def by simp have "?sigma ! (H+i) = \\<^sub>1 (t - 1) ! i" using that zeta0_def zeta1_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 3: "reseq ?sigma w\<^sub>2 = \\<^sub>2 (t - 1)" (is "?l = ?r") proof (rule nth_equalityI) show len: "length ?l = length ?r" using zeta2_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1)) @ \\<^sub>2 (t - 1) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho'_def by simp have "?sigma ! (2*H+i) = \\<^sub>2 (t - 1) ! i" using len that zeta0_def zeta1_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 4: "reseq ?sigma w\<^sub>3 = \\<^sub>0 t" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta0_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1)) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho'_def by simp have "?sigma ! (Z+i) = \\<^sub>0 t ! i" using that Z_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 5: "reseq ?sigma w\<^sub>4 = \\<^sub>1 t" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta1_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 t) @ \\<^sub>1 t @ \\<^sub>2 t @ \ (inputpos m t)" using rho'_def by simp have "?sigma ! (Z+H+i) = \\<^sub>1 t ! i" using that Z_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 6: "reseq ?sigma w\<^sub>5 = \\<^sub>2 t" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using zeta2_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 t @ \\<^sub>1 t) @ \\<^sub>2 t @ \ (inputpos m t)" using rho'_def by simp have "?sigma ! (Z+2*H+i) = \\<^sub>2 t ! i" using that zeta0_def zeta1_def zeta2_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed have 7: "reseq ?sigma w\<^sub>6 = (\ (inputpos m t))" (is "?l = ?r") proof (rule nth_equalityI) show "length ?l = length ?r" using gamma_def by simp show "?l ! i = ?r ! i" if "i < length ?l" for i proof - have 1: "?sigma = (\\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1) @ \\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t) @ \ (inputpos m t) @ []" using rho'_def by simp have "?sigma ! (2*Z+i) = \ (inputpos m t) ! i" using that Z_def gamma_def by (intro nth_append3[OF 1]) auto then show ?thesis using reseq_def that by simp qed qed show ?thesis using ** 1 2 3 4 5 6 7 by simp qed section \The CNF formula $\Phi$\ text \ We can now formulate all the parts $\Phi_0, \dots, \Phi_9$ of the complete formula $\Phi$, and thus $\Phi$ itself. Representing the snapshot in step~0: \ definition PHI0 :: formula ("\\<^sub>0") where "\\<^sub>0 \ \ (\\<^sub>0 0) 1 @ \ (\\<^sub>1 0) 1 @ \ (\\<^sub>2 0) 0" text \The start symbol at the beginning of the input tape:\ definition PHI1 :: formula ("\\<^sub>1") where "\\<^sub>1 \ \ (\ 0) 1" text \The separator \textbf{11} between $x$ and $u$:\ definition PHI2 :: formula ("\\<^sub>2") where "\\<^sub>2 \ \ (\ (2*n+1)) 3 @ \ (\ (2*n+2)) 3" text \The zeros before the symbols of $x$:\ definition PHI3 :: formula ("\\<^sub>3") where "\\<^sub>3 \ concat (map (\i. \ (\ (2*i+1)) 2) [0..The zeros before the symbols of $u$:\ definition PHI4 :: formula ("\\<^sub>4") where "\\<^sub>4 \ concat (map (\i. \ (\ (2*n+2+2*i+1)) 2) [0..

The blank symbols after the input $\langle x, u\rangle$:\ definition PHI5 :: formula ("\\<^sub>5") where "\\<^sub>5 \ concat (map (\i. \ (\ (2*n + 2*p n + 3 + i)) 0) [0..The symbols of $x$:\ definition PHI6 :: formula ("\\<^sub>6") where "\\<^sub>6 \ concat (map (\i. \ (\ (2*i+2)) (if x ! i then 3 else 2)) [0..Constraining the symbols of $u$ to be from $\{\mathbf{0}, \mathbf{1}\}$:\ definition PHI7 :: formula ("\\<^sub>7") where "\\<^sub>7 \ concat (map (\i. \ (\ (2*n+4+2*i))) [0..

Reading a \textbf{1} in the final step to signal acceptance of $\langle x, u\rangle$:\ definition PHI8 :: formula ("\\<^sub>8") where "\\<^sub>8 \ \ (\\<^sub>1 T') 3" text \The snapshots after the first and before the last:\ definition PHI9 :: formula ("\\<^sub>9") where "\\<^sub>9 \ concat (map (\t. \ (Suc t)) [0..The complete formula:\ definition PHI :: formula ("\") where "\ \ \\<^sub>0 @ \\<^sub>1 @ \\<^sub>2 @ \\<^sub>3 @ \\<^sub>4 @ \\<^sub>5 @ \\<^sub>6 @ \\<^sub>7 @ \\<^sub>8 @ \\<^sub>9" section \Correctness of the formula\ text \ We have to show that the formula $\Phi$ is satisfiable if and only if $x \in L$. There is a subsection for both of the implications. Instead of $x \in L$ we will use the right-hand side of the following equivalence. \ lemma L_iff_ex_u: "x \ L \ (\u. length u = p n \ exc \x; u\ T' <.> 1 = \)" proof - have "x \ L \ (\u. length u = p (length x) \ exc \x; u\ (T (length \x; u\)) <.> 1 = \)" using cert by simp also have "... \ (\u. length u = p (length x) \ exc \x; u\ (TT (length \x; u\)) <.> 1 = \)" proof - have "bit_symbols \x; u\" for u by simp then have "exc \x; u\ (TT (length \x; u\)) = exc \x; u\ (T (length \x; u\))" for u using exc_TT_eq_exc_T by blast then show ?thesis by simp qed also have "... \ (\u. length u = p n \ exc \x; u\ T' <.> 1 = \)" using T'_def length_pair m_def by auto finally show ?thesis . qed subsection \$\Phi$ satisfiable implies $x\in L$\ text \ The proof starts from an assignment $\alpha$ satisfying $\Phi$ and shows that there is a string $u$ of length $p(n)$ such that $M$, on input $\langle x, u\rangle$, halts with the output tape head on the symbol \textbf{1}. The overarching idea is that $\alpha$, by satisfying $\Phi$, encodes a string $u$ and a computation of $M$ on $u$ that results in $M$ halting with the output tape head on the symbol~\textbf{1}. The assignment $\alpha$ is an infinite bit string, whose first $N = m'\cdot H$ bits are supposed to encode the first $m'$ symbols on $M$'s input tape, which contains the pair $\langle x, u\rangle$. The first step of the proof is thus to extract a $u$ of length $p(n)$ from the first $N$ bits of $\alpha$. The Formula $\Phi_7$ ensures that the symbols representing $u$ are \textbf{0} or \textbf{1} and thus represent a bit string. Next the proof shows that the first $N$ bits of $\alpha$ encode the relevant portion $y(u)$ of the input tape for the $u$ just extracted, that is, $y(u)_i = \alpha(\gamma_i)$ for $i < m'$. The proof exploits the constraints set by $\Phi_1$ to $\Phi_6$. In particular this implies that $y(u)_{\inputpos(t)} = \alpha(\gamma_{\inputpos(t)})$ for all $t$. The next $Z\cdot (T' + 1)$ bits of $\alpha$ encode $T' + 1$ snapshots. More precisely, we prove $z_0(u, t) = \alpha(\zeta_0^t)$ and $z_1(u, t) = \alpha(\zeta_1^t)$ and $z_2(u, t) = \alpha(\zeta_2^t)$ for all $t \leq T'$. This works by induction on $t$. The case $t = 0$ is covered by the formula $\Phi_0$. For $0 < t \leq T'$ the formulas $\chi_t$ are responsible, which make up $\Phi_9$. Basically $\chi_t$ represents the recursive characterization of the snapshot $z_t$ in terms of earlier snapshots (of $t-1$ and possibly $prev(t)$). This is the trickiest part and we need some preliminary lemmas for that. Once that is done, we know that some bits of $\alpha$, namely $\alpha(\zeta_1(T'))$, encode the symbol under the output tape head after $T'$ steps, that is, when $M$ has halted. Formula $\Phi_8$ ensures that this symbol is \textbf{1}, which concludes the proof. \null \ lemma sat_PHI_imp_ex_u: assumes "satisfiable \" shows "\u. length u = p n \ exc \x; u\ T' <.> 1 = \" proof - obtain \ where \: "\ \ \" using assms satisfiable_def by auto define us where "us = map (\i. unary \ (\ (2*n+4+2*i))) [0..

us ! i = 3" if "i < p n" for i proof - have "\ \ \\<^sub>7" using PHI_def satisfies_def \ by simp then have "\ \ \ (\ (2*n+4+2*i))" using that PHI7_def satisfies_concat_map by auto then have "unary \ (\ (2*n+4+2*i)) = 2 \ unary \ (\ (2*n+4+2*i)) = 3" using Upsilon_unary length_gamma H_gr_2 by simp then show ?thesis using us_def that by simp qed define u :: string where "u = symbols_to_string us" have len_u: "length u = p n" using u_def len_us by simp have "ysymbols u ! i = unary \ (\ i)" if "i < m'" for i proof - consider "i = 0" | "1 \ i \ i < 2 * n + 1" | "2 * n + 1 \ i \ i < 2 * n + 3" | "2 * n + 3 \ i \ i < m + 1" | "i \ m + 1 \ i < m + 1 + T'" using `i < m'` m'_def m_def by linarith then show ?thesis proof (cases) case 1 then have "\ \ \ (\ i) 1" using PHI_def PHI1_def \ satisfies_append by metis then have "unary \ (\ i) = 1" using Psi_unary H_gr_2 gamma_def by simp moreover have "ysymbols u ! i = 1" using 1 by (simp add: ysymbols_def) ultimately show ?thesis by simp next case 2 define j where "j = (i - 1) div 2" then have "j < n" using 2 by auto have "i = 2 * j + 1 \ i = 2 * j + 2" using 2 j_def by auto then show ?thesis proof assume i: "i = 2 * j + 1" have "\ \ \\<^sub>3" using PHI_def satisfies_def \ by simp then have "\ \ \ (\ (2*j+1)) 2" using PHI3_def satisfies_concat_map[OF _ `j < n`] by auto then have "unary \ (\ (2*j+1)) = 2" using Psi_unary H_gr_2 gamma_def by simp moreover have "ysymbols u ! (2*j+1) = 2" using ysymbols_first_at[OF `j < n`] by simp ultimately show ?thesis using i by simp next assume i: "i = 2 * j + 2" have "\ \ \\<^sub>6" using PHI_def satisfies_def \ by simp then have "\ \ \ (\ (2*j+2)) (if x ! j then 3 else 2)" using PHI6_def satisfies_concat_map[OF _ `j < n`] by fastforce then have "unary \ (\ (2*j+2)) = (if x ! j then 3 else 2)" using Psi_unary H_gr_2 gamma_def by simp moreover have "ysymbols u ! (2*j+2) = (if x ! j then 3 else 2)" using ysymbols_first_at[OF `j < n`] by simp ultimately show ?thesis using i by simp qed next case 3 then have "i = 2*n + 1 \ i = 2*n + 2" by auto then show ?thesis proof assume i: "i = 2 * n + 1" then have "\ \ \ (\ i) 3" using PHI_def PHI2_def \ satisfies_append by metis then have "unary \ (\ i) = 3" using Psi_unary H_gr_2 gamma_def by simp moreover have "ysymbols u ! i = 3" using i ysymbols_at_2n1 by simp ultimately show ?thesis by simp next assume i: "i = 2 * n + 2" then have "\ \ \ (\ i) 3" using PHI_def PHI2_def \ satisfies_append by metis then have "unary \ (\ i) = 3" using Psi_unary H_gr_2 gamma_def by simp moreover have "ysymbols u ! i = 3" using i ysymbols_at_2n2 by simp ultimately show ?thesis by simp qed next case 4 moreover define j where "j = (i - 2*n-3) div 2" ultimately have j: "j < p n" using m_def by auto have "i = 2*n+2+2 * j + 1 \ i = 2*n+2+2 * j + 2" using 4 j_def by auto then show ?thesis proof assume i: "i = 2 * n + 2 + 2 * j + 1" have "\ \ \\<^sub>4" using PHI_def satisfies_def \ by simp then have "\ \ \ (\ (2*n+2+2*j+1)) 2" using PHI4_def satisfies_concat_map[OF _ `j < p n`] by auto then have "unary \ (\ (2*n+2+2*j+1)) = 2" using Psi_unary H_gr_2 gamma_def by simp moreover have "ysymbols u ! (2*n+2+2*j+1) = 2" using \j < p n\ ysymbols_second_at(1) len_u by presburger ultimately show ?thesis using i by simp next assume i: "i = 2 * n + 2 + 2 * j + 2" have "us ! j = unary \ (\ (2*n+4+2*j))" using us_def `j < p n` by simp then have "us ! j = unary \ (\ (2*n+2+2*j+2))" by (simp add: numeral_Bit0) then have "unary \ (\ (2*n+2+2*j+2)) = (if u ! j then 3 else 2)" using u_def us23 \j < p local.n\ len_us by fastforce then have *: "unary \ (\ i) = (if u ! j then 3 else 2)" using i by simp have "ysymbols u ! (2*n+2+2*j+2) = (if u ! j then 3 else 2)" using ysymbols_second_at(2) `j < p n` len_u by simp then have "ysymbols u ! i = (if u ! j then 3 else 2)" using i by simp then show ?thesis using * i by simp qed next case 5 then have "\ \ \\<^sub>5" using PHI_def satisfies_def \ by simp then have "\ \ \ (\ (2*n+2*p n + 3 + i')) 0" if "i' < T'" for i' unfolding PHI5_def using \ satisfies_concat_map[OF _ that, of \] that by auto moreover obtain i' where i': "i' < T'" "i = m + 1 + i'" using 5 by (metis add_less_cancel_left le_Suc_ex) ultimately have "\ \ \ (\ i) 0" using m_def numeral_3_eq_3 by simp then have "unary \ (\ i) = 0" using Psi_unary H_gr_2 gamma_def by simp moreover have "ysymbols u ! i = 0" using 5 ysymbols_last len_u by simp ultimately show ?thesis by simp qed qed then have ysymbols: "ysymbols u ! (inputpos m t) = unary \ (\ (inputpos m t))" for t using inputpos_less' len_u by simp have "z0 u t = unary \ (\\<^sub>0 t) \ z1 u t = unary \ (\\<^sub>1 t) \ z2 u t = unary \ (\\<^sub>2 t)" if "t \ T'" for t using that proof (induction t rule: nat_less_induct) case (1 t) show ?case proof (cases "t = 0") case True have "\ \ \\<^sub>0" using \ PHI_def satisfies_def by simp then have 1: "\ \ \ (\\<^sub>0 0) 1" and 2: "\ \ \ (\\<^sub>1 0) 1" and 3: "\ \ \ (\\<^sub>2 0) 0" using PHI0_def by (metis satisfies_append(1), metis satisfies_append, metis satisfies_append(2)) have "unary \ (\\<^sub>0 0) = 1" using Psi_unary[OF _ 1] H_gr_2 by simp moreover have "unary \ (\\<^sub>1 0) = 1" using Psi_unary[OF _ 2] H_gr_2 by simp moreover have "unary \ (\\<^sub>2 0) = 0" using Psi_unary[OF _ 3] by simp moreover have "z0 u 0 = \" using z0_def start_config2 by (simp add: start_config_pos) moreover have "z1 u 0 = \" using z1_def by (simp add: start_config2 start_config_pos) moreover have "z2 u 0 = \" using z2_def by (simp add: start_config_def) ultimately show ?thesis using True by simp next case False then have "t > 0" by simp let ?t = "t - 1" have sat_chi: "\ \ \ t" proof - have "\ \ \\<^sub>9" using \ PHI_def satisfies_def by simp moreover have "?t < T'" using 1 `t > 0` by simp ultimately have "\ \ \ (Suc ?t)" using satisfies_concat_map PHI9_def by auto then show ?thesis using \t > 0\ by simp qed show ?thesis proof (cases "prev m t < t") case True then show ?thesis using satisfies_chi_less z0 z1 z2' 1 len_u ysymbols sat_chi True `t \ T'` len_u by simp next case False then have eq: "prev m t = t" using prev_eq prev_less by blast show ?thesis using satisfies_chi_eq z0 z1' z2' ysymbols sat_chi eq `t > 0` `t \ T'` len_u 1 `t > 0` by simp qed qed qed then have "z1 u T' = unary \ (\\<^sub>1 T')" by simp moreover have "unary \ (\\<^sub>1 T') = 3" proof - have "\ \ \\<^sub>8" using \ PHI_def satisfies_def by simp then have "\ \ \ (\\<^sub>1 T') 3" using PHI8_def by simp then show ?thesis using Psi_unary[of 3 "\\<^sub>1 T'"] length_zeta1 H_gr_2 by simp qed ultimately have "z1 u T' = \" by simp then have "exc \x; u\ T' <.> 1 = \" using z1_def by simp then show ?thesis using len_u by auto qed subsection \$x\in L$ implies $\Phi$ satisfiable\ text \ For the other direction, we assume a string $x \in L$ and show that the formula $\Phi$ derived from it is satisfiable. From $x \in L$ it follows that there is a certificate $u$ of length $p(n)$ such that $M$ on input $\langle x, u\rangle$ halts after $T'$ steps with the output tape head on the symbol~\textbf{1}. An assignment that satisfies $\Phi$ must have its first $N = m' \cdot H$ bits set in such a way that they encode the relevant portion $y(u)$ of the input tape, that is, with $\langle x, u\rangle$ followed by $T'$ blanks. This will take care of satisfying $\Phi_1, \dots, \Phi_7$. The next $Z\cdot (T' + 1)$ bits of $\alpha$ must be set such that they encode the $T' + 1$ snapshots of $M$ when started on $\langle x, u\rangle$. This way $\Phi_0$ and $\Phi_9$ will be satisfied. Finally, $\Phi_8$ is satisfied because by the choice of $u$ the last snapshot contains a \textbf{1} as the symbol under the output tape head. The following function maps a string $u$ to an assignment as just described. \ definition beta :: "string \ assignment" ("\") where "\ u i \ if i < N then let j = i div H; k = i mod H in if j = 0 then k < 1 else if j = 2 * n + 1 \ j = 2 * n + 2 then k < 3 else if j \ 2 * n + 2 * p n + 3 then k < 0 else if odd j then k < 2 else if j \ 2 * n then k < (if x ! (j div 2 - 1) then 3 else 2) else k < (if u ! (j div 2 - n - 2) then 3 else 2) else if i < N + Z * (Suc T') then let t = (i - N) div Z; k = (i - N) mod Z in if k < H then k < z0 u t else if k < 2 * H then k - H < z1 u t else k - 2 * H < z2 u t else False" text \ In order to show that $\beta(u)$ satisfies $\Phi$, we show that it satisfies all parts of $\Phi$. These parts consist mostly of $\Psi$ formulas, whose satisfiability can be proved using lemma~@{thm [source] satisfies_Psi}. To apply this lemma, the following ones will be helpful. \ lemma blocky_gammaI: assumes "\k. k < H \ \ (j * H + k) = (k < l)" shows "blocky \ (\ j) l" unfolding blocky_def gamma_def using assms by simp lemma blocky_zeta0I: assumes "\k. k < H \ \ (N + t * Z + k) = (k < l)" shows "blocky \ (\\<^sub>0 t) l" unfolding blocky_def zeta0_def using assms by simp lemma blocky_zeta1I: assumes "\k. k < H \ \ (N + t * Z + H + k) = (k < l)" shows "blocky \ (\\<^sub>1 t) l" unfolding blocky_def zeta1_def using assms by simp lemma blocky_zeta2I: assumes "\k. k < H \ \ (N + t * Z + 2 * H + k) = (k < l)" shows "blocky \ (\\<^sub>2 t) l" unfolding blocky_def zeta2_def using Z_def assms by simp lemma beta_1: "blocky (\ u) (\ 0) 1" proof (intro blocky_gammaI) fix k :: nat assume k: "k < H" let ?i = "0 * H + k" have "?i < N" using N_eq add_mult_distrib2 k by auto then show "\ u ?i = (k < 1)" using beta_def k by simp qed lemma beta_2a: "blocky (\ u) (\ (2*n+1)) 3" proof (intro blocky_gammaI) fix k :: nat assume k: "k < H" let ?i = "(2*n+1) * H + k" let ?j = "?i div H" let ?k = "?i mod H" have "?i < N" using N_eq add_mult_distrib2 k by auto moreover have j: "?j = 2 * n + 1" using k by (metis add_cancel_left_right div_less div_mult_self3 less_nat_zero_code) moreover have "?k = k" using k by (metis mod_if mod_mult_self3) moreover have "\ ?j = 0" using j by linarith ultimately show "\ u ?i = (k < 3)" using beta_def by simp qed lemma beta_2b: "blocky (\ u) (\ (2*n+2)) 3" proof (intro blocky_gammaI) fix k :: nat assume k: "k < H" let ?i = "(2*n+2) * H + k" let ?j = "?i div H" let ?k = "?i mod H" have "?i < N" using N_eq add_mult_distrib2 k by auto moreover have "?j = 2 * n + 2" using k by (metis add_cancel_left_right div_less div_mult_self3 less_nat_zero_code) moreover have "?k = k" using k by (metis mod_if mod_mult_self3) moreover have "\ ?j = 0" using calculation(2) by linarith ultimately show "\ u ?i = (k < 3)" using beta_def Let_def k by presburger qed lemma beta_3: assumes "ii < n" shows "blocky (\ u) (\ (2 * ii + 1)) 2" proof (intro blocky_gammaI) fix k ::nat assume k: "k < H" let ?i = "(2*ii+1) * H + k" let ?j = "?i div H" let ?k = "?i mod H" have "?i < N" proof - have "?i < (2*n+1) * H + k" using assms k by simp also have "... < (2*n+1) * H + H" using k by simp also have "... = H * (2*n+2)" by simp also have "... \ H * (2*n+3)" by (metis add.commute add.left_commute add_2_eq_Suc le_add2 mult_le_mono2 numeral_3_eq_3) also have "... \ H * (2*n+2*p n+3)" by simp also have "... \ H * (2*n+2*p n+3 + T')" by simp finally have "?i < H * (2*n+2*p n+3 + T')" . then show ?thesis using N_eq by simp qed moreover have j: "?j = 2 * ii + 1" using k by (metis add_cancel_left_right div_less div_mult_self3 less_nat_zero_code) moreover have "?k = k" using k by (metis mod_if mod_mult_self3) moreover have "\ ?j = 0" using j by linarith moreover have "\ (?j = 2*n+1 \ ?j = 2*n+2)" using j assms by simp moreover have "\ ?j \ 2*n+2*p n + 3" using j assms by simp moreover have "odd ?j" using j by simp ultimately show "\ u ?i = (k < 2)" using beta_def by simp qed lemma beta_4: assumes "ii < p n" and "length u = p n" shows "blocky (\ u) (\ (2*n+2+2*ii+1)) 2" proof (intro blocky_gammaI) fix k :: nat assume k: "k < H" let ?i = "(2*n+2+2*ii+1) * H + k" let ?j = "?i div H" let ?k = "?i mod H" have "?i < N" proof - have "?i = (2*n+2*ii+3) * H + k" by (simp add: numeral_3_eq_3) also have "... < (2*n+2*ii+3) * H + H" using k by simp also have "... = H * (2*n+2*ii+4)" by algebra also have "... \ H * (2*n+2*p n+3)" using assms(1) by simp also have "... \ H * (2*n+2*p n+3 + T')" by simp finally have "?i < H * (2*n+2*p n+3 + T')" . then show ?thesis using N_eq by simp qed moreover have j: "?j = 2 * n + 2 + 2 * ii + 1" using k by (metis add_cancel_left_right div_less div_mult_self3 less_nat_zero_code) moreover have "?k = k" using k by (metis mod_if mod_mult_self3) moreover have "\ ?j = 0" using j by linarith moreover have "\ (?j = 2*n+1 \ ?j = 2*n+2)" using j assms by simp moreover have "\ ?j \ 2*n+2*p n + 3" using j assms by simp moreover have "odd ?j" using j by simp ultimately show "\ u ?i = (k < 2)" using beta_def by simp qed lemma beta_5: assumes "ii < T'" shows "blocky (\ u) (\ (2*n+2*p n + 3 + ii)) 0" proof (intro blocky_gammaI) fix k :: nat assume k: "k < H" let ?i = "(2*n+2*p n + 3 + ii) * H + k" let ?j = "?i div H" let ?k = "?i mod H" have "?i < N" proof - have "?i < (2*n+2*p n + 3 + ii) * H + H" using k by simp also have "... \ (2*n+2*p n + 3 + T' - 1) * H + H" proof - have "2*n+2*p n + 3 + ii \ 2*n+2*p n + 3 + T' - 1" using assms by simp then show ?thesis using add_le_mono1 mult_le_mono1 by presburger qed also have "... \ (2*n+2*p n + 2 + T') * H + H" by simp also have "... \ H * (2*n+2*p n + 3 + T')" by (simp add: numeral_3_eq_3) finally have "?i < H * (2*n+2*p n + 3 + T')" . then show ?thesis using N_eq by simp qed moreover have j: "?j = 2 * n + 2*p n + 3 + ii" using k by (metis add_cancel_left_right div_less div_mult_self3 less_nat_zero_code) moreover have "?k = k" using k by (metis mod_if mod_mult_self3) moreover have "\ ?j = 0" using j by linarith moreover have "\ (?j = 2*n+1 \ ?j = 2*n+2)" using j by simp ultimately show "\ u ?i = (k < 0)" using beta_def Let_def k by simp qed lemma beta_6: assumes "ii < n" shows "blocky (\ u) (\ (2 * ii + 2)) (if x ! ii then 3 else 2)" proof (intro blocky_gammaI) fix k :: nat assume k: "k < H" let ?i = "(2*ii+2) * H + k" let ?j = "?i div H" let ?k = "?i mod H" have "?i < N" proof - have "?i \ (2*n+2) * H + k" using assms by simp also have "... < (2*n+2) * H + H" using k by simp also have "... = (2*n+3) * H" by algebra also have "... \ (2*n + 2*p n + 3) * H" by simp also have "... \ (2*n + 2*p n + 3 + T') * H" by simp finally have "?i < (2*n + 2*p n + 3 + T') * H" . then show ?thesis using N_eq by (metis mult.commute) qed moreover have j: "?j = 2 * ii + 2" using k by (metis add_cancel_left_right div_less div_mult_self3 less_nat_zero_code) moreover have "?k = k" using k by (metis mod_if mod_mult_self3) moreover have "\ ?j = 0" using j by linarith moreover have "\ (?j = 2*n+1 \ ?j = 2*n+2)" using j assms by simp moreover have "\ ?j = 2*n+2*p n + 3" using j assms by simp moreover have "\ odd ?j" using j by simp moreover have "?j \ 2 * n" using j assms by simp ultimately show "\ u ?i = (k < (if x ! ii then 3 else 2))" using beta_def by simp qed lemma beta_7: assumes "ii < p n" and "length u = p n" shows "blocky (\ u) (\ (2 * n + 4 + 2 * ii)) (if u ! ii then 3 else 2)" proof (intro blocky_gammaI) fix k :: nat assume k: "k < H" let ?i = "(2*n+4+2*ii) * H + k" let ?j = "?i div H" let ?k = "?i mod H" have "?i < N" proof - have 1: "ii \ p n - 1" using assms(1) by simp have 2: "p n > 0" using assms(1) by simp have "?i \ (2*n+4+2*(p n - 1)) * H + k" using 1 by simp also have "... = (2*n+4+2*p n-2) * H + k" using 2 diff_mult_distrib2 by auto also have "... = (2*n+2+2*p n) * H + k" by simp also have "... < (2*n+2+2*p n) * H + H" using k by simp also have "... = (2*n+3+2*p n) * H" by algebra also have "... = H * (2*n + 2*p n + 3)" by simp also have "... \ H * (2*n + 2*p n + 3 + T')" by simp finally have "?i < H * (2*n + 2*p n + 3 + T')" . then show ?thesis using N_eq by simp qed moreover have j: "?j = 2 * n + 4 + 2 * ii" using k by (metis add_cancel_left_right div_less div_mult_self3 less_nat_zero_code) moreover have "?k = k" using k by (metis mod_if mod_mult_self3) moreover have "\ ?j = 0" using j by linarith moreover have "\ (?j = 2*n+1 \ ?j = 2*n+2)" using j assms by simp moreover have "\ odd ?j" using j by simp moreover have "\ ?j \ 2 * n" using j assms by simp moreover have "?j \ 2 * n + 2 * p n + 2" using j assms by simp ultimately show "\ u ?i = (k < (if u ! ii then 3 else 2))" using beta_def by simp qed lemma beta_zeta0: assumes "t \ T'" shows "blocky (\ u) (\\<^sub>0 t) (z0 u t)" proof (intro blocky_zeta0I) fix k ::nat assume k: "k < H" let ?i = "N + t * Z + k" let ?t = "(?i - N) div Z" let ?k = "(?i - N) mod Z" have "\ ?i < N" by simp moreover have "?i < N + Z * (Suc T')" proof - have "?i \ N + T' * Z + k" using assms by simp also have "... < N + T' * Z + H" using k by simp also have "... \ N + T' * Z + Z" using Z_def by simp also have "... = N + Z * (Suc T')" by simp finally show ?thesis by simp qed moreover have kk: "?k = k" using k Z_def by simp moreover have "?t = t" using k Z_def by simp moreover have "?k < H" using kk k by simp ultimately show "\ u ?i = (k < z0 u t)" using beta_def by simp qed lemma beta_zeta1: assumes "t \ T'" shows "blocky (\ u) (\\<^sub>1 t) (z1 u t)" proof (intro blocky_zeta1I) fix k :: nat assume k: "k < H" let ?i = "N + t * Z + H + k" let ?t = "(?i - N) div Z" let ?k = "(?i - N) mod Z" have "\ ?i < N" by simp moreover have "?i < N + Z * (Suc T')" proof - have "?i \ N + T' * Z + H + k" using assms by simp also have "... < N + T' * Z + H + H" using k by simp also have "... \ N + T' * Z + Z" using Z_def by simp also have "... = N + Z * (Suc T')" by simp finally show ?thesis by simp qed moreover have "?t = t" using k Z_def by simp moreover have kk: "?k = H + k" using k Z_def by simp moreover have "\ ?k < H" using kk by simp moreover have "?k < 2 * H" using kk k by simp ultimately have "\ u ?i = (?k - H < z1 u t)" using beta_def by simp then show "\ u ?i = (k < z1 u t)" using kk by simp qed lemma beta_zeta2: assumes "t \ T'" shows "blocky (\ u) (\\<^sub>2 t) (z2 u t)" proof (intro blocky_zeta2I) fix k :: nat assume k: "k < H" let ?i = "N + t * Z + 2 * H + k" let ?t = "(?i - N) div Z" let ?k = "(?i - N) mod Z" have 1: "2 * H + k < Z" using k Z_def by simp have "\ ?i < N" by simp moreover have "?i < N + Z * (Suc T')" proof - have "?i \ N + T' * Z + 2 * H + k" using assms by simp also have "... < N + T' * Z + 2 * H + H" using k by simp also have "... \ N + T' * Z + Z" using Z_def by simp also have "... = N + Z * (Suc T')" by simp finally show ?thesis by simp qed moreover have "?t = t" using 1 by simp moreover have kk: "?k = 2 * H + k" using k Z_def by simp moreover have "\ ?k < H" using kk by simp moreover have "\ ?k < 2 * H" using kk by simp ultimately have "\ u ?i = (?k - 2 * H < z2 u t)" using beta_def by simp then show "\ u ?i = (k < z2 u t)" using kk by simp qed text \ We can finally show that $\beta(u)$ satisfies $\Phi$ if $u$ is a certificate for $x$. \ lemma satisfies_beta_PHI: assumes "length u = p n" and "exc \x; u\ T' <.> 1 = \" shows "\ u \ \" proof - have "\ u \ \\<^sub>0" proof - have "blocky (\ u) (\\<^sub>0 0) (z0 u 0)" using beta_zeta0 by simp then have "blocky (\ u) (\\<^sub>0 0) 1" using z0_def start_config2 start_config_pos by auto then have "\ u \ \ (\\<^sub>0 0) 1" using satisfies_Psi H_gr_2 by simp moreover have "\ u \ \ (\\<^sub>1 0) 1" proof - have "blocky (\ u) (\\<^sub>1 0) (z1 u 0)" using beta_zeta1 by simp then have "blocky (\ u) (\\<^sub>1 0) 1" using z1_def start_config2 start_config_pos by simp then show ?thesis using satisfies_Psi H_gr_2 by simp qed moreover have "\ u \ \ (\\<^sub>2 0) 0" proof - have "blocky (\ u) (\\<^sub>2 0) (z2 u 0)" using beta_zeta2 by simp then have "blocky (\ u) (\\<^sub>2 0) 0" using z2_def start_config_def by simp then show ?thesis using satisfies_Psi H_gr_2 by simp qed ultimately show ?thesis using PHI0_def satisfies_def by auto qed moreover have "\ u \ \\<^sub>1" using PHI1_def H_gr_2 satisfies_Psi beta_1 by simp moreover have "\ u \ \\<^sub>2" proof - have "\ u \ \ (\ (2*n+1)) 3" using satisfies_Psi H_gr_2 beta_2a by simp moreover have "\ u \ \ (\ (2*n+2)) 3" using satisfies_Psi H_gr_2 beta_2b by simp ultimately show ?thesis using PHI2_def satisfies_def by auto qed moreover have "\ u \ \\<^sub>3" proof - have "\ u \ \ (\ (2*i+1)) 2" if "i < n" for i using satisfies_Psi that H_gr_2 length_gamma less_imp_le_nat beta_3 by simp then show ?thesis using PHI3_def satisfies_concat_map' by simp qed moreover have "\ u \ \\<^sub>4" proof - have "\ u \ \ (\ (2*n+2+2*i+1)) 2" if "i < p n" for i using satisfies_Psi that H_gr_2 length_gamma less_imp_le_nat beta_4 assms(1) by simp then show ?thesis using PHI4_def satisfies_concat_map' by simp qed moreover have "\ u \ \\<^sub>5" proof - have "\ u \ \ (\ (2*n+2*p n+3+i)) 0" if "i < T'" for i using satisfies_Psi that H_gr_2 length_gamma less_imp_le_nat beta_5 assms(1) by simp then show ?thesis using PHI5_def satisfies_concat_map' by simp qed moreover have "\ u \ \\<^sub>6" proof - have "\ u \ \ (\ (2*i+2)) (if x ! i then 3 else 2)" if "i < n" for i using satisfies_Psi that H_gr_2 length_gamma less_imp_le_nat beta_6 by simp then show ?thesis using PHI6_def satisfies_concat_map' by simp qed moreover have "\ u \ \\<^sub>7" proof - have "\ u \ \ (\ (2*n+4+2*i))" if "i < p n" for i proof - have "blocky (\ u) (\ (2*n+4+2*i)) 2 \ blocky (\ u) (\ (2*n+4+2*i)) 3" using assms that beta_7[of i u] by (metis (full_types)) then have "\ u \ \ (\ (2*n+4+2*i)) 2 \ \ u \ \ (\ (2*n+4+2*i)) 3" using satisfies_Psi H_gr_2 by auto then show ?thesis using Psi_2_imp_Upsilon Psi_3_imp_Upsilon H_gr_2 length_gamma by auto qed then show ?thesis using PHI7_def satisfies_concat_map' by simp qed moreover have "\ u \ \\<^sub>8" using PHI8_def H_gr_2 assms satisfies_Psi z1_def beta_zeta1 by (metis One_nat_def Suc_1 Suc_leI length_zeta1 nat_le_linear numeral_3_eq_3) moreover have "\ u \ \\<^sub>9" proof - have *: "unary (\ u) (\ i) = ysymbols u ! i" if "i < length (ysymbols u)" for i proof - have "i < m'" using assms length_ysymbols that by simp then consider "i = 0" | "1 \ i \ i < 2*n + 1" | "2*n+1 \ i \ i < 2*n+3" | "2*n+3 \ i \ i < m+1" | "i \ m + 1 \ i < m + 1 + T'" using m'_def m_def by linarith then show ?thesis proof (cases) case 1 then show ?thesis using ysymbols_at_0 blocky_imp_unary H_gr_2 beta_1 by simp next case 2 moreover define j where "j = (i - 1) div 2" ultimately have j: "j < n" "i = 2 * j + 1 \ i = 2 * j + 2" by auto show ?thesis proof (cases "i = 2 * j + 1") case True then show ?thesis using ysymbols_first_at(1) blocky_imp_unary H_gr_2 j(1) beta_3 by simp next case False then have "i = 2 * j + 2" using j(2) by simp then show ?thesis using ysymbols_first_at(2) blocky_imp_unary H_gr_2 j(1) beta_4 beta_6 by simp qed next case 3 show ?thesis proof (cases "i = 2*n+1") case True then show ?thesis using ysymbols_at_2n1 blocky_imp_unary H_gr_2 beta_2a by simp next case False then have "i = 2*n+2" using 3 by simp then show ?thesis using ysymbols_at_2n2 blocky_imp_unary H_gr_2 beta_2b by simp qed next case 4 moreover define j where "j = (i - 2*n-3) div 2" ultimately have j: "j < p n" "i = 2*n+2+2 * j + 1 \ i = 2*n+2+2 * j + 2" using j_def m_def by auto show ?thesis proof (cases "i = 2*n+2+2 * j + 1") case True then show ?thesis using ysymbols_second_at(1) assms(1) blocky_imp_unary H_gr_2 j(1) beta_4 by simp next case False then have i: "i = 2*n+4+2 * j" using j(2) by simp then have "ysymbols u ! (2*n+2+2*j+2) = (if u ! j then 3 else 2)" using ysymbols_second_at(2) assms j(1) by simp then have "ysymbols u ! (2*n+4+2*j) = (if u ! j then 3 else 2)" by (metis False i j(2)) then have "ysymbols u ! i = (if u ! j then 3 else 2)" using i by simp then show ?thesis using beta_7[OF j(1)] blocky_imp_unary H_gr_2 length_gamma i assms(1) by simp qed next case 5 then obtain ii where ii: "ii < T'" "i = m + 1 + ii" by (metis le_iff_add nat_add_left_cancel_less) have "blocky (\ u) (\ (2*n+2*p n + 3 + ii)) 0" using beta_5[OF ii(1)] by simp then have "blocky (\ u) (\ i) 0" using ii(2) m_def numeral_3_eq_3 by simp then have "unary (\ u) (\ i) = 0" using blocky_imp_unary by simp moreover have "ysymbols u ! i = 0" using ysymbols_last[OF assms(1)] 5 by simp ultimately show ?thesis by simp qed qed have "\ u \ \ (Suc t)" (is "\ u \ \ ?t") if "t < T'" for t proof (cases "prev m ?t < ?t") case True have t: "?t \ T'" using that by simp then have "unary (\ u) (\\<^sub>0 ?t) = z0 u ?t" using blocky_imp_unary z0_le beta_zeta0 by simp moreover have "ysymbols u ! (inputpos m ?t) = unary (\ u) (\ (inputpos m ?t))" using * assms(1) inputpos_less' length_ysymbols by simp ultimately have "unary (\ u) (\\<^sub>0 ?t) = unary (\ u) (\ (inputpos m ?t))" using assms(1) z0 by simp moreover have "unary (\ u) (\\<^sub>1 ?t) = z1 u ?t" using beta_zeta1 blocky_imp_unary z1_le t by simp moreover have "unary (\ u) (\\<^sub>2 ?t) = z2 u ?t" using beta_zeta2 blocky_imp_unary z2_le t by simp moreover have "unary (\ u) (\\<^sub>0 (prev m ?t)) = z0 u (prev m ?t)" using beta_zeta0 blocky_imp_unary z0_le t True by simp moreover have "unary (\ u) (\\<^sub>1 (prev m ?t)) = z1 u (prev m ?t)" using beta_zeta1 blocky_imp_unary z1_le t True by simp moreover have "unary (\ u) (\\<^sub>2 (prev m ?t)) = z2 u (prev m ?t)" using beta_zeta2 blocky_imp_unary z2_le t True by simp moreover have "unary (\ u) (\\<^sub>0 (?t - 1)) = z0 u (?t - 1)" using beta_zeta0 blocky_imp_unary z0_le t True by simp moreover have "unary (\ u) (\\<^sub>1 (?t - 1)) = z1 u (?t - 1)" using beta_zeta1 blocky_imp_unary z1_le t True by simp moreover have "unary (\ u) (\\<^sub>2 (?t - 1)) = z2 u (?t - 1)" using beta_zeta2 blocky_imp_unary z2_le t True by simp ultimately show ?thesis using True assms(1) satisfies_chi_less[OF True] t z1 z2' by (metis bot_nat_0.extremum less_nat_zero_code nat_less_le) next case False then have prev: "prev m ?t = ?t" using prev_le by (meson le_neq_implies_less) have t: "?t \ T'" using that by simp then have "unary (\ u) (\\<^sub>0 ?t) = z0 u ?t" using beta_zeta0 blocky_imp_unary z0_le by simp moreover have "ysymbols u ! (inputpos m ?t) = unary (\ u) (\ (inputpos m ?t))" using * assms(1) inputpos_less' length_ysymbols by simp ultimately have "unary (\ u) (\\<^sub>0 ?t) = unary (\ u) (\ (inputpos m ?t))" using assms(1) z0 by simp moreover have "unary (\ u) (\\<^sub>1 ?t) = z1 u ?t" using beta_zeta1 blocky_imp_unary z1_le t by simp moreover have "z1 u ?t = \" using z1' beta_zeta1 assms(1) prev t by simp moreover have "unary (\ u) (\\<^sub>2 ?t) = z2 u ?t" using beta_zeta2 blocky_imp_unary z2_le t by simp moreover have "unary (\ u) (\\<^sub>0 (?t - 1)) = z0 u (?t - 1)" using beta_zeta0 blocky_imp_unary z0_le t by simp moreover have "unary (\ u) (\\<^sub>1 (?t - 1)) = z1 u (?t - 1)" using beta_zeta1 blocky_imp_unary z1_le t by simp moreover have "unary (\ u) (\\<^sub>2 (?t - 1)) = z2 u (?t - 1)" using beta_zeta2 blocky_imp_unary z2_le t by simp ultimately show ?thesis using satisfies_chi_eq[OF prev] start_config2 start_config_pos t that z1_def z2 assms(1) by (metis (no_types, lifting) One_nat_def Suc_1 Suc_less_eq add_diff_inverse_nat execute.simps(1) less_one n_not_Suc_n plus_1_eq_Suc) qed then show ?thesis using PHI9_def satisfies_concat_map' by presburger qed ultimately show ?thesis using satisfies_append' PHI_def by simp qed corollary ex_u_imp_sat_PHI: assumes "length u = p n" and "exc \x; u\ T' <.> 1 = \" shows "satisfiable \" using satisfies_beta_PHI assms satisfiable_def by auto text \ The formula $\Phi$ has the desired property: \ theorem L_iff_satisfiable: "x \ L \ satisfiable \" using L_iff_ex_u ex_u_imp_sat_PHI sat_PHI_imp_ex_u by auto end (* locale reduction_sat_x *) end diff --git a/thys/Cook_Levin/Sat_TM_CNF.thy b/thys/Cook_Levin/Sat_TM_CNF.thy --- a/thys/Cook_Levin/Sat_TM_CNF.thy +++ b/thys/Cook_Levin/Sat_TM_CNF.thy @@ -1,4579 +1,4579 @@ section \Turing machines for the parts of $\Phi$\label{s:tmcnf}\ theory Sat_TM_CNF imports Aux_TM_Reducing begin text \ In this section we build Turing machines for all parts $\Phi_0,\dots,\Phi_9$ of the CNF formula $\Phi$. Some of them ($\Phi_0$, $\Phi_1$, $\Phi_2$, and $\Phi_8$) are just fixed-length sequences of $\Psi$ formulas constructible by fixed-length sequences of @{const tm_Psigamma} machines. Others ($\Phi_3$, $\Phi_4$, $\Phi_5$, $\Phi_6$) are variable-length and require looping over a @{const tm_Psigamma} machine. The TM for $\Phi_7$ is a loop over @{const tm_Upsilongamma}. Lastly, the TM for $\Phi_9$ is a loop over a TM that generates the formulas $\chi_t$. Ideally we would want to prove the semantics of the TMs inside the locale @{locale reduction_sat}, in which $\Phi$ was defined. However, we use locales to prove the semantics of TMs, and locales cannot be nested. For this reason we have to define the TMs on the theory level and prove their semantics there, too, just as we have done with all TMs until now. In the next chapter the semantics lemmas will be transferred to the locale @{locale reduction_sat}. Unlike most TMs so far, the TMs in this section are not meant to be reusable but serve a special purpose, namely to be combined into one large TM computing $\Phi$. For this reason the TMs are somewhat peculiar. For example, they write their output to the fixed tape~$1$ rather than having a parameter for the output tape. They also often expect the tapes to be initialized in a very special way. Moreover, the TMs often leave the work tapes in a ``dirty'' state with remnants of intermediate calculations. The combined TM for all of $\Phi$ will simply allocate a new batch of work tapes for every individual TM. \ subsection \A Turing machine for $\Phi_0$\ text \ The next Turing machine expects a number $i$ on tape $j$ and a number $H$ on tape $j + 1$ and outputs to tape $1$ the formula $\Psi([i\dots H, (i+1)\dots H), 1) \land \Psi([(i+1)\dots H, (i+2)\dots H), 1) \land \Psi([(i+2)\dots H, (i+3)\dots H), 0)$, which is just $\Phi_0$ for suitable values of $i$ and $H$. \ definition tm_PHI0 :: "tapeidx \ machine" where "tm_PHI0 j \ tm_setn (j + 2) 1 ;; tm_Psigamma j ;; tm_extendl_erase 1 (j + 6) ;; tm_incr j ;; tm_Psigamma j ;; tm_extendl_erase 1 (j + 6) ;; tm_incr j ;; tm_setn (j + 2) 0 ;; tm_Psigamma j ;; tm_extendl 1 (j + 6)" lemma tm_PHI0_tm: assumes "0 < j" and "j + 8 < k" and "G \ 6" shows "turing_machine k G (tm_PHI0 j)" unfolding tm_PHI0_def using assms tm_Psigamma_tm tm_extendl_tm tm_erase_cr_tm tm_times2_tm tm_incr_tm tm_setn_tm tm_cr_tm tm_extendl_erase_tm by simp locale turing_machine_PHI0 = fixes j :: tapeidx begin definition "tm1 \ tm_setn (j + 2) 1" definition "tm2 \ tm1 ;; tm_Psigamma j" definition "tm3 \ tm2 ;; tm_extendl_erase 1 (j + 6)" definition "tm5 \ tm3 ;; tm_incr j" definition "tm6 \ tm5 ;; tm_Psigamma j" definition "tm7 \ tm6 ;; tm_extendl_erase 1 (j + 6)" definition "tm9 \ tm7 ;; tm_incr j" definition "tm10 \ tm9 ;; tm_setn (j + 2) 0" definition "tm11 \ tm10 ;; tm_Psigamma j" definition "tm12 \ tm11 ;; tm_extendl 1 (j + 6)" lemma tm12_eq_tm_PHI0: "tm12 = tm_PHI0 j" using tm12_def tm11_def tm10_def tm9_def tm7_def tm6_def tm5_def using tm3_def tm2_def tm1_def tm_PHI0_def by simp context fixes tps0 :: "tape list" and k idx H :: nat assumes jk: "length tps0 = k" "1 < j" "j + 8 < k" and H: "H \ 3" assumes tps0: "tps0 ! 1 = (\[]\, 1)" "tps0 ! j = (\idx\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\[]\, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\[]\, 1)" "tps0 ! (j + 7) = (\[]\, 1)" "tps0 ! (j + 8) = (\[]\, 1)" begin definition "tps1 \ tps0 [j + 2 := (\1\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 12" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps0 tps1_def jk) show "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" using tps0 jk canrepr_0 by simp show "ttt = 10 + 2 * nlength 0 + 2 * nlength 1" using assms canrepr_1 by simp qed definition "tps2 \ tps0 [j + 2 := (\1\\<^sub>N, 1), j + 6 := (\nll_Psi (idx * H) H 1\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 12 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: assms tps0 H tps1_def tps2_def jk) definition "tps3 \ tps0 [j + 2 := (\1\\<^sub>N, 1), 1 := nlltape (nll_Psi (idx * H) H (Suc 0)), j + 6 := (\[]\, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 23 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H (Suc 0))" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps0 H tps2_def tps3_def jk time: assms) show "tps2 ! 1 = nlltape []" using tps2_def jk nllcontents_Nil tps0 by simp show "tps3 = tps2 [1 := nlltape ([] @ nll_Psi (idx * H) H (Suc 0)), j + 6 := (\[]\, 1)]" unfolding tps3_def tps2_def using jk by (simp add: list_update_swap) qed definition "tps5 \ tps0 [j + 2 := (\1\\<^sub>N, 1), 1 := nlltape (nll_Psi (idx * H) H (Suc 0)), j + 6 := (\[]\, 1), j := (\Suc idx\\<^sub>N, 1)]" lemma tm5 [transforms_intros]: assumes "ttt = 28 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def by (tform tps: tps0 H tps3_def tps5_def jk time: assms) definition "tps6 \ tps0 [j := (\Suc idx\\<^sub>N, 1), j + 2 := (\1\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc idx * H) H (Suc 0)\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nll_Psi (idx * H) H 1)]" lemma tm6 [transforms_intros]: assumes "ttt = 28 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def proof (tform tps: tps0 H tps5_def tps6_def jk time: assms) show "tps6 = tps5[j + 6 := (\nll_Psi (Suc idx * H) H (Suc 0)\\<^sub>N\<^sub>L\<^sub>L, 1)]" unfolding tps5_def tps6_def using jk by (simp add: list_update_swap[of j] list_update_swap[of _ "j + 6"]) qed definition "tps7 \ tps0 [j := (\Suc idx\\<^sub>N, 1), j + 2 := (\1\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nll_Psi (idx * H) H 1 @ nll_Psi (H + idx * H) H 1)]" lemma tm7 [transforms_intros]: assumes "ttt = 39 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1)" shows "transforms tm7 tps0 ttt tps7" unfolding tm7_def by (tform tps: assms tps6_def tps7_def jk) definition "tps9 \ tps0 [j := (\Suc (Suc idx)\\<^sub>N, 1), j + 2 := (\1\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nll_Psi (idx * H) H 1 @ nll_Psi (H + idx * H) H 1)]" lemma tm9 [transforms_intros]: assumes "ttt = 44 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1) + 2 * nlength (Suc idx)" shows "transforms tm9 tps0 ttt tps9" unfolding tm9_def proof (tform tps: tps0 H tps7_def tps9_def jk time: assms) show "tps9 = tps7[j := (\Suc (Suc idx)\\<^sub>N, 1)]" using tps9_def tps7_def jk by (simp add: list_update_swap) qed definition "tps10 \ tps0 [j := (\Suc (Suc idx)\\<^sub>N, 1), j + 2 := (\0\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nll_Psi (idx * H) H 1 @ nll_Psi (H + idx * H) H 1)]" lemma tm10 [transforms_intros]: assumes "ttt = 56 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1) + 2 * nlength (Suc idx)" shows "transforms tm10 tps0 ttt tps10" unfolding tm10_def proof (tform tps: tps0 H tps9_def tps10_def jk) show "ttt = 44 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1) + 2 * nlength (Suc idx) + (10 + 2 * nlength (Suc 0) + 2 * nlength 0) " using assms canrepr_1 by simp qed definition "tps11 \ tps0 [j := (\Suc (Suc idx)\\<^sub>N, 1), j + 2 := (\0\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc (Suc idx) * H) H 0\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nll_Psi (idx * H) H 1 @ nll_Psi (H + idx * H) H 1)]" lemma tm11 [transforms_intros]: assumes "ttt = 56 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1) + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2" shows "transforms tm11 tps0 ttt tps11" unfolding tm11_def by (tform tps: tps0 H tps10_def tps11_def jk time: assms) definition "tps12 \ tps0 [j := (\Suc (Suc idx)\\<^sub>N, 1), j + 2 := (\0\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc (Suc idx) * H) H 0\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nll_Psi (idx * H) H 1 @ nll_Psi (H + idx * H) H 1 @ nll_Psi (Suc (Suc idx) * H) H 0)]" lemma tm12: assumes "ttt = 60 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1) + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 2 * nlllength (nll_Psi (H + (H + idx * H)) H 0)" shows "transforms tm12 tps0 ttt tps12" unfolding tm12_def by (tform tps: tps11_def tps12_def jk assms) lemma tm12': assumes "ttt = 5627 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2" shows "transforms tm12 tps0 ttt tps12" proof - let ?ttt = "60 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * nlllength (nll_Psi (idx * H) H 1) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1) + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 2 * nlllength (nll_Psi (H + (H + idx * H)) H 0)" have "?ttt \ 60 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * nlllength (nll_Psi (H + idx * H) H 1) + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 2 * nlllength (nll_Psi (H + (H + idx * H)) H 0)" using nlllength_nll_Psi_le'[of "idx * H" "2 * H + idx * H" "H"] by simp also have "... \ 60 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 2 * nlllength (nll_Psi (H + (H + idx * H)) H 0)" using nlllength_nll_Psi_le'[of "H + idx * H" "2 * H + idx * H" "H"] by simp also have "... \ 60 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 4 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 4 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 2 * H * (3 + nlength (3 * H + idx * H))" using nlllength_nll_Psi_le'[of "H + (H + idx * H)" "2 * H + idx * H" "H"] by simp also have "... = 60 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 10 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2" by simp also have "... \ 60 + 1851 * H ^ 4 * (nlength (Suc (Suc idx)))\<^sup>2 + 10 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2" using nlength_mono linear_le_pow by simp also have "... \ 60 + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 10 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 2 * nlength (Suc idx) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2" using nlength_mono linear_le_pow by simp also have "... = 60 + 5553 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 10 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength idx + 2 * nlength (Suc idx)" by simp also have "... \ 60 + 5553 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 10 * H * (3 + nlength (3 * H + idx * H)) + 2 * nlength (Suc idx) + 2 * nlength (Suc idx)" using nlength_mono by simp also have "... = 60 + 5553 * H ^ 4 * (nlength (Suc (Suc (Suc idx))))\<^sup>2 + 10 * H * (3 + nlength (3 * H + idx * H)) + 4 * nlength (Suc idx)" by simp also have "... \ 60 + 5553 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2 + 10 * H * (3 + nlength (3 * H + idx * H)) + 4 * nlength (Suc idx)" proof - have "Suc (Suc (Suc idx)) \ 3 * H + idx * H" proof (cases "idx = 0") case True then show ?thesis using H by simp next case False then show ?thesis using H by (metis One_nat_def Suc3_eq_add_3 comm_semiring_class.distrib le_Suc_eq less_eq_nat.simps(1) mult.commute mult_1 mult_le_mono1 nle_le not_numeral_le_zero) qed then have "nlength (Suc (Suc (Suc idx))) \ 3 + nlength (3 * H + idx * H)" using nlength_mono trans_le_add2 by presburger then have "nlength (Suc (Suc (Suc idx))) ^ 2 \ (3 + nlength (3 * H + idx * H)) ^ 2" by simp then show ?thesis by simp qed also have "... \ 60 + 5553 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2 + 10 * H ^ 4 * (3 + nlength (3 * H + idx * H)) + 4 * nlength (Suc idx)" using linear_le_pow by simp also have "... \ 60 + 5553 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2 + 10 * H ^ 4 * (3 + nlength (3 * H + idx * H)) ^ 2 + 4 * nlength (Suc idx)" using linear_le_pow by simp also have "... = 60 + 5563 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2 + 4 * nlength (Suc idx)" by simp also have "... \ 60 + 5563 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2 + 4 * H^4 * (3 + nlength (3 * H + idx * H))^2" proof - have "idx \ idx * H" using H by simp then have "Suc idx \ 3 * H + idx * H" using H by linarith then have "nlength (Suc idx) \ 3 + nlength (3 * H + idx * H)" using nlength_mono trans_le_add2 by presburger also have "... \ (3 + nlength (3 * H + idx * H)) ^ 2" by (simp add: power2_eq_square) also have "... \ H * (3 + nlength (3 * H + idx * H)) ^ 2" using H by simp also have "... \ H^4 * (3 + nlength (3 * H + idx * H)) ^ 2" using linear_le_pow by simp finally have "nlength (Suc idx) \ H^4 * (3 + nlength (3 * H + idx * H)) ^ 2" . then show ?thesis by simp qed also have "... = 60 + 5567 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2" by simp also have "... \ 60 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2 + 5567 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2" using H linear_le_pow by simp also have "... = 5627 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2" by simp finally have "?ttt \ ttt" using assms by simp then show ?thesis using tm12 transforms_monotone by simp qed end (* context tps0 *) end (* locale turing_machine_PHI0 *) lemma transforms_tm_PHI0I: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k idx H :: nat assumes "length tps = k" and "1 < j" and "j + 8 < k" and "H \ 3" assumes "tps ! 1 = (\[]\, 1)" "tps ! j = (\idx\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\[]\, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\[]\, 1)" "tps ! (j + 7) = (\[]\, 1)" "tps ! (j + 8) = (\[]\, 1)" assumes "tps' = tps [j := (\Suc (Suc idx)\\<^sub>N, 1), j + 2 := (\0\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc (Suc idx) * H) H 0\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nll_Psi (idx * H) H 1 @ nll_Psi (H + idx * H) H 1 @ nll_Psi (Suc (Suc idx) * H) H 0)]" assumes "ttt = 5627 * H ^ 4 * (3 + nlength (3 * H + idx * H))\<^sup>2" shows "transforms (tm_PHI0 j) tps ttt tps'" proof - interpret loc: turing_machine_PHI0 j . show ?thesis using loc.tps12_def loc.tm12' loc.tm12_eq_tm_PHI0 assms by metis qed subsection \A Turing machine for $\Phi_1$\ text \ The next TM expects a number $H$ on tape $j + 1$ and appends to the formula on tape $1$ the formula $\Psi([0, H), 1)$. \ definition tm_PHI1 :: "tapeidx \ machine" where "tm_PHI1 j \ tm_setn (j + 2) 1 ;; tm_Psigamma j ;; tm_extendl 1 (j + 6)" lemma tm_PHI1_tm: assumes "0 < j" and "j + 7 < k" and "G \ 6" shows "turing_machine k G (tm_PHI1 j)" unfolding tm_PHI1_def using assms tm_Psigamma_tm tm_setn_tm tm_extendl_tm by simp locale turing_machine_PHI1 = fixes j :: tapeidx begin definition "tm1 \ tm_setn (j + 2) 1" definition "tm2 \ tm1 ;; tm_Psigamma j" definition "tm3 \ tm2 ;; tm_extendl 1 (j + 6)" lemma tm3_eq_tm_PHI1: "tm3 = tm_PHI1 j" using tm3_def tm2_def tm1_def tm_PHI1_def by simp context fixes tps0 :: "tape list" and k idx H :: nat and nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j" "j + 7 < k" and H: "H \ 3" assumes tps0: "tps0 ! 1 = nlltape nss" "tps0 ! j = (\0\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\[]\, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\[]\, 1)" "tps0 ! (j + 7) = (\[]\, 1)" begin definition "tps1 \ tps0 [j + 2 := (\1\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 12" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps0 tps1_def jk) show "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" using tps0 jk canrepr_0 by simp show "ttt = 10 + 2 * nlength 0 + 2 * nlength 1" using assms canrepr_1 by simp qed definition "tps2 \ tps0 [j + 2 := (\1\\<^sub>N, 1), j + 6 := (\nll_Psi 0 H 1\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 12 + 1851 * H ^ 4" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps0 H tps1_def tps2_def jk) show "ttt = 12 + 1851 * H ^ 4 * (nlength (Suc 0))\<^sup>2" using canrepr_1 assms by simp qed definition "tps3 \ tps0 [j + 2 := (\1\\<^sub>N, 1), j + 6 := (\nll_Psi 0 H 1\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nss @ nll_Psi 0 H 1)]" lemma tm3: assumes "ttt = 16 + 1851 * H ^ 4 + 2 * nlllength (nll_Psi 0 H 1)" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def by (tform tps: tps0 H tps2_def tps3_def jk time: assms) lemma tm3': assumes "ttt = 1875 * H ^ 4" shows "transforms tm3 tps0 ttt tps3" proof - let ?ttt = "16 + 1851 * H ^ 4 + 2 * nlllength (nll_Psi 0 H 1)" have "?ttt \ 16 + 1851 * H ^ 4 + 2 * H * (3 + nlength H)" using nlllength_nll_Psi_le by (metis (mono_tags, lifting) add_left_mono mult.assoc nat_mult_le_cancel1 plus_nat.add_0 rel_simps(51)) also have "... = 16 + 1851 * H ^ 4 + 6 * H + 2 * H * nlength H" by algebra also have "... \ 16 + 1851 * H ^ 4 + 6 * H + 2 * H * H" using nlength_le_n by simp also have "... \ 16 + 1851 * H ^ 4 + 6 * H * H + 2 * H * H" by simp also have "... = 16 + 1851 * H ^ 4 + 8 * H^2" by algebra also have "... \ 16 + 1851 * H ^ 4 + 8 * H^4" using pow_mono'[of 2 4 H] by simp also have "... = 16 + 1859 * H ^ 4" by simp also have "... \ 16 * H^4 + 1859 * H ^ 4" using H by simp also have "... = 1875 * H ^ 4" by simp finally have "?ttt \ 1875 * H ^ 4" . then show ?thesis using assms tm3 transforms_monotone by simp qed end (* context tps0 *) end (* locale turing_machine_PHI1 *) lemma transforms_tm_PHI1I: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k H :: nat and nss :: "nat list list" assumes "length tps = k" and "1 < j" and "j + 7 < k" and "H \ 3" assumes "tps ! 1 = nlltape nss" "tps ! j = (\0\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\[]\, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\[]\, 1)" "tps ! (j + 7) = (\[]\, 1)" assumes "tps' = tps [j + 2 := (\1\\<^sub>N, 1), j + 6 := (\nll_Psi 0 H 1\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nss @ nll_Psi 0 H 1)]" assumes "ttt = 1875 * H ^ 4" shows "transforms (tm_PHI1 j) tps ttt tps'" proof - interpret loc: turing_machine_PHI1 j . show ?thesis using loc.tps3_def loc.tm3' loc.tm3_eq_tm_PHI1 assms by metis qed subsection \A Turing machine for $\Phi_2$\ text \ The next TM expects a number $i$ on tape $j$ and a number $H$ on tape $j + 1$. It appends to the formula on tape~$1$ the formula $\Psi([(2i+1)H, (2i+2)H), 3) \land \Psi([(2i+2)H, (2i+3)H), 3)$. \ definition tm_PHI2 :: "tapeidx \ machine" where "tm_PHI2 j \ tm_times2 j ;; tm_incr j ;; tm_setn (j + 2) 3 ;; tm_Psigamma j ;; tm_extendl_erase 1 (j + 6) ;; tm_incr j ;; tm_Psigamma j ;; tm_extendl 1 (j + 6)" lemma tm_PHI2_tm: assumes "0 < j" and "j + 8 < k" and "G \ 6" shows "turing_machine k G (tm_PHI2 j)" unfolding tm_PHI2_def using assms tm_Psigamma_tm tm_extendl_tm tm_erase_cr_tm tm_times2_tm tm_incr_tm tm_setn_tm tm_cr_tm tm_extendl_erase_tm by simp locale turing_machine_PHI2 = fixes j :: tapeidx begin definition "tm1 \ tm_times2 j" definition "tm2 \ tm1 ;; tm_incr j" definition "tm3 \ tm2 ;; tm_setn (j + 2) 3" definition "tm4 \ tm3 ;; tm_Psigamma j" definition "tm5 \ tm4 ;; tm_extendl_erase 1 (j + 6)" definition "tm7 \ tm5 ;; tm_incr j" definition "tm8 \ tm7 ;; tm_Psigamma j" definition "tm9 \ tm8 ;; tm_extendl 1 (j + 6)" lemma tm9_eq_tm_PHI2: "tm9 = tm_PHI2 j" using tm9_def tm8_def tm7_def tm5_def tm4_def tm3_def tm2_def tm1_def tm_PHI2_def by simp context fixes tps0 :: "tape list" and k idx H :: nat and nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j" "j + 7 < k" and H: "H \ 3" assumes tps0: "tps0 ! 1 = nlltape nss" "tps0 ! j = (\idx\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\[]\, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\[]\, 1)" "tps0 ! (j + 7) = (\[]\, 1)" begin definition "tps1 \ tps0 [j := (\2 * idx\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 5 + 2 * nlength idx" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def by (tform tps: tps0 tps1_def jk assms) definition "tps2 \ tps0 [j := (\2 * idx + 1\\<^sub>N, 1)]" lemma tm2: assumes "ttt = 10 + 2 * nlength idx + 2 * nlength (2 * idx)" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: tps0 H tps1_def tps2_def jk assms) lemma tm2' [transforms_intros]: assumes "ttt = 12 + 4 * nlength idx" shows "transforms tm2 tps0 ttt tps2" proof - have "10 + 2 * nlength idx + 2 * nlength (2 * idx) \ 10 + 2 * nlength idx + 2 * (Suc (nlength idx))" using nlength_times2 by (meson add_left_mono mult_le_mono2) then have "10 + 2 * nlength idx + 2 * nlength (2 * idx) \ ttt" using assms by simp then show ?thesis using tm2 transforms_monotone by simp qed definition "tps3 \ tps0 [j := (\2 * idx + 1\\<^sub>N, 1), j + 2 := (\3\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 26 + 4 * nlength idx" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps0 H tps2_def tps3_def jk) show "tps2 ! (j + 2) = (\0\\<^sub>N, 1)" using tps2_def jk canrepr_0 tps0 by simp show "ttt = 12 + 4 * nlength idx + (10 + 2 * nlength 0 + 2 * nlength 3)" using nlength_3 assms by simp qed definition "tps4 \ tps0 [j := (\2 * idx + 1\\<^sub>N, 1), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc (2 * idx) * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tm4 [transforms_intros]: assumes "ttt = 26 + 4 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def by (tform tps: tps0 H tps3_def tps4_def jk time: assms) definition "tps5 \ tps0 [j := (\2 * idx + 1\\<^sub>N, 1), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nss @ nll_Psi (H + 2 * idx * H) H 3)]" lemma tm5 [transforms_intros]: assumes "ttt = 37 + 4 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3)" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def by (tform tps: tps0 H tps4_def tps5_def jk time: assms) definition "tps7 \ tps0 [j := (\2 * idx + 2\\<^sub>N, 1), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nss @ nll_Psi (H + 2 * idx * H) H 3)]" lemma tm7: assumes "ttt = 42 + 4 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3) + 2 * nlength (Suc (2 * idx))" shows "transforms tm7 tps0 ttt tps7" unfolding tm7_def proof (tform tps: tps0 H tps5_def tps7_def jk time: assms) show "tps7 = tps5[j := (\Suc (Suc (2 * idx))\\<^sub>N, 1)]" using tps5_def tps7_def jk by (simp add: list_update_swap) qed lemma tm7' [transforms_intros]: assumes "ttt = 44 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3)" shows "transforms tm7 tps0 ttt tps7" proof - let ?ttt = "42 + 4 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3) + 2 * nlength (Suc (2 * idx))" have "?ttt \ 42 + 4 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3) + 2 * (Suc (nlength idx))" using nlength_times2plus1 by (metis add.commute add_left_mono mult_le_mono2 plus_1_eq_Suc) then have "?ttt \ ttt" using assms by simp then show ?thesis using tm7 transforms_monotone by simp qed definition "tps8 \ tps0 [j := (\2 * idx + 2\\<^sub>N, 1), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc (Suc (2 * idx)) * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nss @ nll_Psi (H + 2 * idx * H) H 3)]" lemma tm8 [transforms_intros]: assumes "ttt = 44 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2" shows "transforms tm8 tps0 ttt tps8" unfolding tm8_def proof (tform tps: tps0 H tps7_def tps8_def jk time: assms) show "tps8 = tps7 [j + 6 := (\nll_Psi (Suc (Suc (2 * idx)) * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1)]" unfolding tps8_def tps7_def by (simp add: list_update_swap[of "j+6"]) qed definition "tps9 \ tps0 [j := (\2 * idx + 2\\<^sub>N, 1), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc (Suc (2 * idx)) * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nss @ nll_Psi (H + 2 * idx * H) H 3 @ nll_Psi (2 * H + 2 * idx * H) H 3)]" lemma tm9: assumes "ttt = 48 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2 + 2 * nlllength (nll_Psi (2 * H + 2 * idx * H) H 3)" shows "transforms tm9 tps0 ttt tps9" unfolding tm9_def by (tform tps: tps0 H tps9_def tps8_def jk time: assms) lemma tm9': assumes "ttt = 3764 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2" shows "transforms tm9 tps0 ttt tps9" proof - let ?ttt = "48 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * nlllength (nll_Psi (H + 2 * idx * H) H 3) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2 + 2 * nlllength (nll_Psi (2 * H + 2 * idx * H) H 3)" have "?ttt \ 48 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 4 * H * (3 + nlength (2 * H + 2 * idx * H + H)) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2 + 2 * nlllength (nll_Psi (2 * H + 2 * idx * H) H 3)" using nlllength_nll_Psi_le'[of "H + 2 * idx * H" "2 * H + 2 * idx * H" H 3] by simp also have "... \ 48 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 5 * H * (3 + nlength (2 * H + 2 * idx * H + H)) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2 + 3 * H * (3 + nlength (2 * H + 2 * idx * H + H))" using nlllength_nll_Psi_le'[of "2 * H + 2 * idx * H" "2 * H + 2 * idx * H" H 3] by simp also have "... = 48 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (2 * idx))))\<^sup>2 + 8 * H * (3 + nlength (2 * H + 2 * idx * H + H)) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2" by simp also have "... \ 48 + 6 * nlength idx + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2 + 8 * H * (3 + nlength (2 * H + 2 * idx * H + H)) + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2" using H4_nlength H by simp also have "... = 48 + 6 * nlength idx + 3702 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * idx)))))\<^sup>2 + 8 * H * (3 + nlength (3 * H + 2 * idx * H))" by simp also have "... \ 48 + 6 * nlength idx + 3702 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2 + 8 * H * (3 + nlength (3 * H + 2 * idx * H))" proof - have "Suc (Suc (Suc (2 * idx))) \ 3 * H + 2 * idx * H" using H by (metis One_nat_def Suc3_eq_add_3 Suc_eq_plus1_left add_leD1 comm_monoid_mult_class.mult_1 distrib_right mult.commute mult_le_mono1) then have "nlength (Suc (Suc (Suc (2 * idx)))) \ 3 + nlength (3 * H + 2 * idx * H)" using nlength_mono trans_le_add2 by blast then show ?thesis by simp qed also have "... \ 48 + 6 * nlength idx + 3702 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2 + 8 * H^4 * (3 + nlength (3 * H + 2 * idx * H))" using linear_le_pow by simp also have "... \ 48 + 6 * nlength idx + 3702 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2 + 8 * H^4 * (3 + nlength (3 * H + 2 * idx * H))^2" using linear_le_pow by simp also have "... = 48 + 6 * nlength idx + 3710 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2" by simp also have "... \ 48 + 3716 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2" proof - have "nlength idx \ nlength (3 * H + 2 * idx * H)" using H by (intro nlength_mono) (simp add: trans_le_add2) also have "... \ 3 + nlength (3 * H + 2 * idx * H)" by simp also have "... \ (3 + nlength (3 * H + 2 * idx * H)) ^ 2" using linear_le_pow by simp also have "... \ H ^ 4 * (3 + nlength (3 * H + 2 * idx * H)) ^ 2" using H by simp finally have "nlength idx \ H ^ 4 * (3 + nlength (3 * H + 2 * idx * H)) ^ 2" . then show ?thesis by simp qed also have "... \ 3764 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2" proof - have "1 \ nlength (3 * H + 2 * idx * H)" using H nlength_0 by (metis One_nat_def Suc_leI add_eq_0_iff_both_eq_0 length_0_conv length_greater_0_conv mult_Suc not_numeral_le_zero numeral_3_eq_3) also have "... \ 3 + nlength (3 * H + 2 * idx * H)" by simp also have "... \ (3 + nlength (3 * H + 2 * idx * H)) ^ 2" using linear_le_pow by simp also have "... \ H ^ 4 * (3 + nlength (3 * H + 2 * idx * H)) ^ 2" using H by simp finally have "1 \ H ^ 4 * (3 + nlength (3 * H + 2 * idx * H)) ^ 2" . then show ?thesis by simp qed finally have "?ttt \ 3764 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2" . then show ?thesis using assms tm9 transforms_monotone by simp qed end (* context tps0 *) end (* locale turing_machine_PHI2 *) lemma transforms_tm_PHI2I: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k idx H :: nat and nss :: "nat list list" assumes "length tps = k" and "1 < j" and "j + 8 < k" and "H \ 3" assumes "tps ! 1 = nlltape nss" "tps ! j = (\idx\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\[]\, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\[]\, 1)" "tps ! (j + 7) = (\[]\, 1)" "tps ! (j + 8) = (\[]\, 1)" assumes "ttt = 3764 * H ^ 4 * (3 + nlength (3 * H + 2 * idx * H))\<^sup>2" assumes "tps' = tps [j := (\2 * idx + 2\\<^sub>N, 1), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\nll_Psi (Suc (Suc (2 * idx)) * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1), 1 := nlltape (nss @ nll_Psi (H + 2 * idx * H) H 3 @ nll_Psi (2 * H + 2 * idx * H) H 3)]" shows "transforms (tm_PHI2 j) tps ttt tps'" proof - interpret loc: turing_machine_PHI2 j . show ?thesis using loc.tm9' loc.tps9_def loc.tm9_eq_tm_PHI2 assms by simp qed subsection \Turing machines for $\Phi_3$, $\Phi_4$, and $\Phi_5$\ text \ The CNF formulas $\Phi_3$, $\Phi_4$, and $\Phi_5$ have a similar structure and can thus be handled by the same Turing machine. The following TM has a parameter $step$ and the usual tape parameter $j$. It expects on tape $j$ a number $idx$, on tape $j + 1$ a number $H$, on tape $j + 2$ a number $\kappa$, and on tape $j + 8$ the number $idx + step \cdot numiter$ for some number $numiter$. It appends to the CNF formula on tape~$1$ the formula $\Psi(\gamma_{idx}, \kappa) \land \Psi(\gamma_{idx + step\cdot (numiter - 1)}, \kappa)$, where $\gamma_i = [iH, (i+1)H)$. \null \ definition tm_PHI345 :: "nat \ tapeidx \ machine" where "tm_PHI345 step j \ WHILE tm_equalsn j (j + 8) (j + 3) ; \rs. rs ! (j + 3) = \ DO tm_Psigamma j ;; tm_extendl_erase 1 (j + 6) ;; tm_plus_const step j DONE" lemma tm_PHI345_tm: assumes "G \ 6" and "0 < j" and "j + 8 < k" shows "turing_machine k G (tm_PHI345 step j)" unfolding tm_PHI345_def using assms tm_equalsn_tm tm_Psigamma_tm tm_extendl_erase_tm tm_plus_const_tm turing_machine_loop_turing_machine by simp locale turing_machine_PHI345 = fixes step :: nat and j :: tapeidx begin definition "tmC \ tm_equalsn j (j + 8) (j + 3)" definition "tm1 \ tm_Psigamma j" definition "tm2 \ tm1 ;; tm_extendl_erase 1 (j + 6)" definition "tm4 \ tm2 ;; tm_plus_const step j" definition "tmL \ WHILE tmC ; \rs. rs ! (j + 3) = \ DO tm4 DONE" lemma tmL_eq_tm_PHI345: "tmL = tm_PHI345 step j" unfolding tmL_def tm4_def tm2_def tm1_def tm_PHI345_def tmC_def by simp context fixes tps0 :: "tape list" and numiter H k idx kappa :: nat and nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j" "j + 8 < k" and H: "H \ 3" and kappa: "kappa \ H" and step: "step > 0" assumes tps0: "tps0 ! 1 = nlltape nss" "tps0 ! j = (\idx\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\kappa\\<^sub>N, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\[]\, 1)" "tps0 ! (j + 7) = (\[]\, 1)" "tps0 ! (j + 8) = (\idx + step * numiter\\<^sub>N, 1)" begin definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j := (\idx + step * t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..idx\\<^sub>N = \idx + step * 0\\<^sub>N" by simp moreover have "nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..<0])) = nlltape nss" using nllcontents_Nil by simp ultimately show ?thesis using tpsL_def tps0 jk by (metis list_update_id) qed definition tpsC :: "nat \ tape list" where "tpsC t \ tps0 [j := (\idx + step * t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..t = numiter\\<^sub>B, 1)]" lemma tmC: assumes "t \ numiter" and "ttt = 3 * nlength (idx + step * t) + 7" shows "transforms tmC (tpsL t) ttt (tpsC t)" unfolding tmC_def proof (tform tps: tps0 tpsL_def jk) show "tpsL t ! (j + 3) = (\0\\<^sub>N, 1)" using canrepr_0 jk tpsL_def tps0 by simp show "(0::nat) \ 1" by simp show "tpsC t = (tpsL t) [j + 3 := (\idx + step * t = idx + step * numiter\\<^sub>B, 1)]" using step tpsC_def jk tpsL_def tps0 by simp show "ttt = 3 * nlength (min (idx + step * t) (idx + step * numiter)) + 7" using assms by (simp add: min_def) qed lemma tmC' [transforms_intros]: assumes "t \ numiter" and "ttt = 3 * nlength (idx + step * numiter) + 7" shows "transforms tmC (tpsL t) ttt (tpsC t)" proof - have "3 * nlength (idx + step * t) + 7 \ ttt" using assms nlength_mono by simp then show ?thesis using assms tmC transforms_monotone by blast qed definition tpsL0 :: "nat \ tape list" where "tpsL0 t \ tps0 [j := (\idx + step * t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0.. tape list" where "tpsL1 t \ tps0 [j := (\idx + step * t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..nll_Psi ((idx+step*t) * H) H kappa\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tm1 [transforms_intros]: assumes "t < numiter" and "ttt = 1851 * H ^ 4 * (nlength (Suc (idx+step*t)))\<^sup>2" shows "transforms tm1 (tpsL0 t) ttt (tpsL1 t)" unfolding tm1_def by (tform tps: H kappa tps0 tpsL0_def tpsL1_def jk time: assms(2)) definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [j := (\idx + step * t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..[]\, 1), 1 := nlltape ((nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..2 + (11 + 4 * nlllength (nll_Psi ((idx + step * t) * H) H kappa))" shows "transforms tm2 (tpsL0 t) ttt (tpsL2 t)" unfolding tm2_def by (tform tps: assms H kappa tps0 tpsL1_def tpsL2_def jk) definition tpsL2' :: "nat \ tape list" where "tpsL2' t \ tps0 [j := (\idx + step * t\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..2 + (11 + 4 * nlllength (nll_Psi ((idx + step * t) * H) H kappa))" shows "transforms tm2 (tpsL0 t) ttt (tpsL2' t)" proof - let ?ttt = "1851 * H ^ 4 * (nlength (Suc (idx + step * t)))\<^sup>2 + (11 + 4 * nlllength (nll_Psi ((idx + step * t) * H) H kappa))" have "?ttt \ 1851 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + (11 + 4 * nlllength (nll_Psi ((idx + step * t) * H) H kappa))" proof - have "Suc (idx + step * t) \ Suc (idx + step * (numiter - 1))" using assms(1) step by simp also have "... = Suc (idx + step * numiter - step)" by (metis Nat.add_diff_assoc One_nat_def Suc_le_eq add_less_same_cancel1 assms(1) mult.right_neutral nat_mult_le_cancel_disj nat_neq_iff not_add_less1 right_diff_distrib') also have "... \ idx + step * numiter" using step Suc_le_eq assms(1) by simp finally have "Suc (idx + step * t) \ idx + step * numiter" . then have "nlength (Suc (idx + step * t)) \ nlength (idx + step * numiter)" using nlength_mono by simp then show ?thesis by simp qed then have "transforms tm2 (tpsL0 t) ttt (tpsL2 t)" using assms tm2 transforms_monotone by blast then show ?thesis using tpsL2' by simp qed definition tpsL2'' :: "nat \ tape list" where "tpsL2'' t \ tps0 [j := (\idx + step * t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..[]\, 1)]" lemma tpsL2'': "tpsL2'' t = tpsL2' t" proof - have "nll_Psi ((idx+step*t) * H) H kappa = nll_Psi (H * (idx+step*t)) H kappa" by (simp add: mult.commute) then have "concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..i. nll_Psi (H * (idx + step * i)) H kappa) [0.. 5 * H ^ 4 * nlength (idx + step * numiter)^2" proof - have "nlllength (nll_Psi ((idx + step * t) * H) H kappa) \ H * (3 + nlength ((idx + step * t) * H + H))" using nlllength_nll_Psi_le by simp also have "... \ H * (3 + nlength ((idx + step * numiter) * H))" proof - have "(idx + 1 + step * t) \ (idx + step * Suc t)" using step by simp then have "(idx + 1 + step * t) \ (idx + step * numiter)" using assms(1) Suc_le_eq by auto then have "(idx + 1 + step * t) * H \ (idx + step * numiter) * H" using mult_le_cancel2 by blast then show ?thesis using nlength_mono by simp qed also have "... = 3 * H + H * nlength ((idx + step * numiter) * H)" by algebra also have "... \ 3 * H + H * (nlength (idx + step * numiter) + nlength H)" using nlength_prod by simp also have "... \ 3 * H + H * (nlength (idx + step * numiter) + H)" using nlength_le_n by simp also have "... = 3 * H + H ^ 2 + H * nlength (idx + step * numiter)" by algebra also have "... \ 3 * H^4 + H ^ 2 + H * nlength (idx + step * numiter)" using linear_le_pow by simp also have "... \ 3 * H^4 + H ^ 4 + H * nlength (idx + step * numiter)" using pow_mono' by simp also have "... \ 4 * H^4 + H ^ 4 * nlength (idx + step * numiter)" using linear_le_pow by simp also have "... \ 4 * H^4 + H ^ 4 * nlength (idx + step * numiter)^2" using linear_le_pow by simp also have "... \ 5 * H ^ 4 * nlength (idx + step * numiter)^2" proof - have "idx + step * numiter > 0" using assms(1) step by simp then have "nlength (idx + step * numiter) > 0" using nlength_0 by simp then have "nlength (idx + step * numiter) ^ 2 > 0" by simp then have "H ^ 4 \ H ^ 4 * nlength (idx + step * numiter) ^ 2" by (metis One_nat_def Suc_leI mult_numeral_1_right nat_mult_le_cancel_disj numerals(1)) then show ?thesis by simp qed finally show ?thesis . qed lemma tm2'' [transforms_intros]: assumes "t < numiter" and "ttt = 1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11" shows "transforms tm2 (tpsL0 t) ttt (tpsL2'' t)" proof - have "transforms tm2 (tpsL0 t) ttt (tpsL2' t)" using tm2'[OF assms(1)] nlllength_nll_Psi[OF assms(1)] transforms_monotone assms(2) by simp then show ?thesis using tpsL2' tpsL2'' by simp qed definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [j := (\idx + step * Suc t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..[]\, 1)]" lemma tm4: assumes "t < numiter" and "ttt = 1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11 + step * (5 + 2 * nlength (idx + step * t + step))" shows "transforms tm4 (tpsL0 t) ttt (tpsL4 t)" unfolding tm4_def proof (tform tps: assms(1) H kappa tps0 tpsL2''_def tpsL4_def jk) have "idx + step * Suc t = idx + step * t + step" by simp then show "tpsL4 t = (tpsL2'' t)[j := (\idx + step * t + step\\<^sub>N, 1)]" unfolding tpsL4_def tpsL2''_def using jk by (simp only: list_update_swap[of _ j] list_update_overwrite) show "ttt = 1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11 + step * (5 + 2 * nlength (idx + step * t + step))" using assms(2) . qed lemma tm4': assumes "t < numiter" and "ttt = (6 * step + 1882) * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2" shows "transforms tm4 (tpsC t) ttt (tpsL4 t)" proof - let ?ttt = "1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11 + step * (5 + 2 * nlength (idx + step * t + step))" have "idx + step * t + step \ idx + step * numiter" using assms(1) by (metis Suc_le_eq add.assoc add.commute add_le_cancel_left add_mult_distrib2 mult_le_mono2 mult_numeral_1_right numerals(1) plus_1_eq_Suc) then have "?ttt \ 1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11 + step * (5 + 2 * nlength (idx + step * numiter))" using nlength_mono by simp also have "... = 1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11 + step * 5 + step * 2 * nlength (idx + step * numiter)" by algebra also have "... \ 1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11 + step * 5 + step * 2 * nlength (idx + step * numiter)^2" by (simp add: linear_le_pow) also have "... \ 1871 * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + 11 + step * 5 + step * H^4 * nlength (idx + step * numiter)^2" proof - have "2 < H" using H by simp then have "2 < H ^ 4" using linear_le_pow by (meson le_trans not_less zero_less_numeral) then show ?thesis by simp qed also have "... = (step + 1871) * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + (11 + step * 5)" by algebra also have "... \ (step + 1871) * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2 + (11 + step * 5) * H ^ 4 * (nlength (idx + step * numiter))^2" proof - have "H ^ 4 * (nlength (idx + step * numiter))^2 > 0" using step assms(1) nlength_0 H by auto then show ?thesis - by (smt (z3) One_nat_def Suc_leI add_mono_thms_linordered_semiring(2) mult.assoc mult_numeral_1_right nat_mult_le_cancel_disj numeral_code(1)) + by (smt (verit) One_nat_def Suc_leI add_mono_thms_linordered_semiring(2) mult.assoc mult_numeral_1_right nat_mult_le_cancel_disj numeral_code(1)) qed also have "... = (6 * step + 1882) * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2" by algebra finally have "?ttt \ (6 * step + 1882) * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2" . then have "transforms tm4 (tpsL0 t) ttt (tpsL4 t)" using tm4 assms transforms_monotone by blast then show ?thesis using tpsL0_eq_tpsC assms(1) by simp qed lemma tm4'' [transforms_intros]: assumes "t < numiter" and "ttt = (6 * step + 1882) * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2" shows "transforms tm4 (tpsC t) ttt (tpsL (Suc t))" proof - have "tpsL4 t = tpsL (Suc t)" using tpsL4_def tpsL_def jk tps0 list_update_id[of tps0 "j + 6"] by (simp add: list_update_swap) then show ?thesis using tm4' assms by metis qed lemma tmL: assumes "ttt = Suc numiter * (9 + (6 * step + 1885) * (H ^ 4 * (nlength (idx + step * numiter))\<^sup>2))" and "nn = numiter" shows "transforms tmL (tpsL 0) ttt (tpsC nn)" unfolding tmL_def proof (tform tps: assms(2)) let ?ttt = "(6 * step + 1882) * H ^ 4 * (nlength (idx + step * numiter))\<^sup>2" show "\t. t < nn \ read (tpsC t) ! (j + 3) = \" using assms(2) tpsC_def jk read_ncontents_eq_0 by simp show "read (tpsC nn) ! (j + 3) \ \" using assms(2) tpsC_def jk read_ncontents_eq_0 by simp show "nn * (3 * nlength (idx + step * numiter) + 7 + ?ttt + 2) + (3 * nlength (idx + step * numiter) + 7) + 1 \ ttt" (is "?lhs \ ttt") proof - let ?g = "H ^ 4 * (nlength (idx + step * numiter))\<^sup>2" have "nlength (idx + step * numiter) \ nlength (idx + step * numiter)^2" using linear_le_pow by simp moreover have "H ^ 4 > 0" using H by simp ultimately have *: "nlength (idx + step * numiter) \ ?g" by (metis ab_semigroup_mult_class.mult_ac(1) le_square mult.left_commute nat_mult_le_cancel_disj power2_eq_square) have "?lhs = numiter * (3 * nlength (idx + step * numiter) + 9 + ?ttt) + 3 * nlength (idx + step * numiter) + 8" using assms(2) by simp also have "... \ numiter * (3 * nlength (idx + step * numiter) + 9 + ?ttt) + 3 * ?g + 8" using * by simp also have "... \ numiter * (3 * ?g + 9 + (6 * step + 1882) * ?g) + 3 * ?g + 8" using * by simp also have "... = numiter * (9 + (6 * step + 1885) * ?g) + 3 * ?g + 8" by algebra also have "... \ numiter * (9 + (6 * step + 1885) * ?g) + (6 * step + 1885) * ?g + 8" by simp also have "... \ numiter * (9 + (6 * step + 1885) * ?g) + (9 + (6 * step + 1885) * ?g)" by simp also have "... = Suc numiter * (9 + (6 * step + 1885) * ?g)" by simp finally show ?thesis using assms(1) by simp qed qed lemma tmL': assumes "ttt = Suc numiter * (9 + (6 * step + 1885) * (H ^ 4 * (nlength (idx + step * numiter))\<^sup>2))" shows "transforms tmL tps0 ttt (tpsC numiter)" using assms tmL tpsL_eq_tps0 by simp end (* context tps0 *) end (* locale turing_machine_PHI345 *) lemma transforms_tm_PHI345I: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt step numiter H k idx kappa :: nat and nss :: "nat list list" assumes "length tps = k" and "1 < j" and "j + 8 < k" and "H \ 3" and "kappa \ H" and "step > 0" assumes "tps ! 1 = nlltape nss" "tps ! j = (\idx\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\kappa\\<^sub>N, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\[]\, 1)" "tps ! (j + 7) = (\[]\, 1)" "tps ! (j + 8) = (\idx + step * numiter\\<^sub>N, 1)" assumes "ttt = Suc numiter * (9 + (6 * step + 1885) * (H ^ 4 * (nlength (idx + step * numiter))\<^sup>2))" assumes "tps' = tps [j := (\idx + step * numiter\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (idx + step * i)) H kappa) [0..1\\<^sub>N, 1)]" shows "transforms (tm_PHI345 step j) tps ttt tps'" proof - interpret loc: turing_machine_PHI345 step j . show ?thesis using assms loc.tmL' loc.tpsC_def loc.tmL_eq_tm_PHI345 by simp qed subsection \A Turing machine for $\Phi_6$\ text \ The next Turing machine expects a symbol sequence $zs$ on input tape~$0$, the number~$2$ on tape $j$, and a number $H$ on tape $j + 1$. It appends to the CNF formula on tape~$1$ the formula $\bigwedge_{i=0}^{|zs| - 1} \Psi(\gamma_{2+2i}, z_i)$, where $z_i$ is $2$ or $3$ if $zs_i$ is \textbf{0} or \textbf{1}, respectively. \ definition tm_PHI6 :: "tapeidx \ machine" where "tm_PHI6 j \ WHILE [] ; \rs. rs ! 0 \ \ DO IF \rs. rs ! 0 = \ THEN tm_setn (j + 2) 2 ELSE tm_setn (j + 2) 3 ENDIF ;; tm_Psigamma j ;; tm_extendl_erase 1 (j + 6) ;; tm_setn (j + 2) 0 ;; tm_right 0 ;; tm_plus_const 2 j DONE" lemma tm_PHI6_tm: assumes "G \ 6" and "0 < j" and "j + 7 < k" shows "turing_machine k G (tm_PHI6 j)" unfolding tm_PHI6_def using assms tm_Psigamma_tm tm_extendl_erase_tm tm_plus_const_tm turing_machine_loop_turing_machine turing_machine_branch_turing_machine tm_right_tm tm_setn_tm Nil_tm by simp locale turing_machine_PHI6 = fixes j :: tapeidx begin definition "tm1 \ IF \rs. rs ! 0 = \ THEN tm_setn (j + 2) 2 ELSE tm_setn (j + 2) 3 ENDIF" definition "tm2 \ tm1 ;; tm_Psigamma j" definition "tm3 \ tm2 ;; tm_extendl_erase 1 (j + 6)" definition "tm4 \ tm3 ;; tm_setn (j + 2) 0" definition "tm5 \ tm4 ;; tm_right 0" definition "tm6 \ tm5 ;; tm_plus_const 2 j" definition "tmL \ WHILE [] ; \rs. rs ! 0 \ \ DO tm6 DONE" lemma tmL_eq_tm_PHI6: "tmL = tm_PHI6 j" using tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def tm_PHI6_def tmL_def by simp context fixes tps0 :: "tape list" and k H :: nat and zs :: "symbol list" and nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j" "j + 7 < k" and H: "H \ 3" and zs: "bit_symbols zs" assumes tps0: "tps0 ! 0 = (\zs\, 1)" "tps0 ! 1 = nlltape nss" "tps0 ! j = (\2\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\[]\, 1)" "tps0 ! (j + 7) = (\[]\, 1)" begin lemma H0: "H > 0" using H by simp lemma H_mult: "x \ H * x" "x \ x * H" using H by simp_all definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [0 := (\zs\, Suc t), j := (\2 + 2 * t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (2 + 2 * i)) H (zs ! i)) [0..zs\, 1) = (\zs\, Suc 0)" by simp moreover have "nlltape (concat (map (\i. nll_Psi (H * (2 + 2 * i)) H (zs ! i)) [0..<0])) = (\[]\, 1)" using nllcontents_Nil by simp moreover have "2 = Suc (Suc 0)" by simp ultimately show ?thesis using tpsL_def tps0 jk by simp (metis One_nat_def Suc_1 list_update_id) qed definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0 [0 := (\zs\, Suc t), j := (\2 + 2 * t\\<^sub>N, 1), j + 2 := (\zs ! t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (2 + 2 * i)) H (zs ! i)) [0.. \ tpsL1 t = (tpsL t)[j + 2 := (\2\\<^sub>N, 1)]" using * tpsL_def tpsL1_def jk by (simp add: list_update_swap) show "tpsL1 t = (tpsL t)[j + 2 := (\3\\<^sub>N, 1)]" if "read (tpsL t) ! \ \ \" proof - have "zs ! t = \" using * that zs assms(2) by auto then show ?thesis using tpsL_def tpsL1_def jk by (simp add: list_update_swap) qed show "10 + 2 * nlength 0 + 2 * nlength 2 + 2 \ ttt" using nlength_0 nlength_2 assms(1) by simp show "10 + 2 * nlength 0 + 2 * nlength 3 + 1 \ ttt" using nlength_0 nlength_3 assms(1) by simp qed definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [0 := (\zs\, Suc t), j := (\2 + 2 * t\\<^sub>N, 1), j + 2 := (\zs ! t\\<^sub>N, 1), j + 6 := (\nll_Psi (2 * H + 2 * t * H) H (zs ! t)\\<^sub>N\<^sub>L\<^sub>L, Suc 0), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (2 + 2 * i)) H (zs ! i)) [0..2" and "t < length zs" shows "transforms tm2 (tpsL t) ttt (tpsL2 t)" unfolding tm2_def proof (tform tps: assms(2) tps0 H tpsL1_def tpsL2_def jk time: assms(1)) show "zs ! t \ H" using assms(2) H zs by auto qed definition tpsL3 :: "nat \ tape list" where "tpsL3 t \ tps0 [0 := (\zs\, Suc t), j := (\2 + 2 * t\\<^sub>N, 1), j + 2 := (\zs ! t\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (2 * H + H * (2 * i)) H (zs ! i)) [0..2 + 4 * nlllength (nll_Psi (2 * H + 2 * t * H) H (zs ! t))" and "t < length zs" shows "transforms tm3 (tpsL t) ttt (tpsL3 t)" unfolding tm3_def proof (tform tps: assms(2) tps0 H tpsL2_def tpsL3_def jk time: assms(1)) have "nll_Psi (2 * H + H * (2 * t)) H = nll_Psi (2 * H + 2 * t * H) H" by (simp add: mult.commute) then have "(nss @ concat (map (\i. nll_Psi (2 * H + H * (2 * i)) H (zs ! i)) [0..i. nll_Psi (2 * H + H * (2 * i)) H (zs ! i)) [0..i. nll_Psi (2 * H + H * (2 * i)) H (zs ! i)) [0..[]\, 1)]" unfolding tpsL3_def tpsL2_def jk by (simp add: list_update_swap) qed definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [0 := (\zs\, Suc t), j := (\2 + 2 * t\\<^sub>N, 1), j + 2 := (\0\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (2 * H + H * (2 * i)) H (zs ! i)) [0..2 + 4 * nlllength (nll_Psi (2 * H + 2 * t * H) H (zs ! t))" and "t < length zs" shows "transforms tm4 (tpsL t) ttt (tpsL4 t)" unfolding tm4_def proof (tform tps: assms(2) tpsL3_def tpsL4_def jk) have "zs ! t = 2 \ zs ! t = 3" using zs assms(2) by auto then show "ttt = 27 + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * t)))))\<^sup>2 + 4 * nlllength (nll_Psi (2 * H + 2 * t * H) H (zs ! t)) + (10 + 2 * nlength (zs ! t) + 2 * nlength 0)" using nlength_2 nlength_3 assms(1) by auto qed definition tpsL5 :: "nat \ tape list" where "tpsL5 t \ tps0 [0 := (\zs\, Suc (Suc t)), j := (\2 + 2 * t\\<^sub>N, 1), j + 2 := (\0\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (2 * H + H * (2 * i)) H (zs ! i)) [0..2 + 4 * nlllength (nll_Psi (2 * H + 2 * t * H) H (zs ! t))" and "t < length zs" shows "transforms tm5 (tpsL t) ttt (tpsL5 t)" unfolding tm5_def proof (tform tps: assms(2) tps0 H tpsL4_def tpsL5_def jk time: assms(1)) have neq: "0 \ j" using jk by simp have "tpsL4 t ! 0 = (\zs\, Suc t)" using tpsL4_def jk by simp then show "tpsL5 t = (tpsL4 t)[0 := tpsL4 t ! 0 |+| 1]" unfolding tpsL5_def tpsL4_def jk by (simp add: list_update_swap[OF neq] list_update_swap[of _ 0]) qed definition tpsL6 :: "nat \ tape list" where "tpsL6 t \ tps0 [0 := (\zs\, Suc (Suc t)), j := (\2 + 2 * (Suc t)\\<^sub>N, 1), j + 2 := (\0\\<^sub>N, 1), j + 6 := (\[]\, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (2 * H + H * (2 * i)) H (zs ! i)) [0..2 + 4 * nlllength (nll_Psi (2 * H + 2 * t * H) H (zs ! t)) + 4 * nlength (4 + 2 * t)" and "t < length zs" shows "transforms tm6 (tpsL t) ttt (tpsL6 t)" unfolding tm6_def proof (tform tps: tps0 H tpsL5_def tpsL6_def jk assms(2)) show "tpsL6 t = (tpsL5 t)[j := (\Suc (Suc (2 * t)) + 2\\<^sub>N, 1)]" using tpsL6_def tpsL5_def jk by (simp add: list_update_swap) have *: "4 + 2 * t = Suc (Suc (Suc (Suc (2 * t))))" by simp then show "ttt = 42 + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * t)))))\<^sup>2 + 4 * nlllength (nll_Psi (2 * H + 2 * t * H) H (zs ! t)) + 2 * (5 + 2 * nlength (Suc (Suc (2 * t)) + 2))" using assms(1) * by simp qed lemma tpsL6_eq_tpsL: "tpsL6 t = tpsL (Suc t)" proof - have "tpsL (Suc t) ! (j + 6) = (\[]\, 1)" using tpsL_def tps0 jk by simp then have "tpsL (Suc t) = (tpsL (Suc t))[j + 6 := (\[]\, 1)]" using list_update_id by metis moreover have "tpsL (Suc t) ! (j + 2) = (\0\\<^sub>N, 1)" using tpsL_def tps0 jk canrepr_0 by simp ultimately have "tpsL (Suc t) = (tpsL (Suc t))[j + 6 := (\[]\, 1), j + 2 := (\0\\<^sub>N, 1)]" using list_update_id by metis moreover have "tpsL6 t = (tpsL (Suc t))[j + 6 := (\[]\, 1), j + 2 := (\0\\<^sub>N, 1)]" unfolding tpsL6_def tpsL_def by (simp add: list_update_swap) ultimately show ?thesis by simp qed lemma tm6': assumes "ttt = 133648 * H ^ 6 * length zs^2" and "t < length zs" shows "transforms tm6 (tpsL t) ttt (tpsL (Suc t))" proof - have **: "Suc (2 * length zs)^2 \ 9 * length zs ^ 2" proof - have "Suc (2 * length zs)^2 = 1 + 2 * 2 * length zs + (2 * length zs)^2" by (metis add.commute add_Suc_shift mult.assoc nat_mult_1_right one_power2 plus_1_eq_Suc power2_sum) also have "... = 1 + 4 * length zs + 4 * length zs^2" by simp also have "... \ length zs + 4 * length zs + 4 * length zs^2" using assms(2) by simp also have "... = 5 * length zs + 4 * length zs^2" by simp also have "... \ 5 * length zs^2 + 4 * length zs^2" using linear_le_pow by simp also have "... = 9 * length zs^2" by simp finally show ?thesis by simp qed have *: "t \ length zs - 1" using assms(2) by simp let ?ttt = "52 + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * t)))))\<^sup>2 + 4 * nlllength (nll_Psi (2 * H + 2 * t * H) H (zs ! t)) + 4 * nlength (4 + 2 * t)" have "?ttt = 52 + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * t)))))\<^sup>2 + 4 * nlllength (nll_Psi (2 * H + H * (2 * t)) H (zs ! t)) + 4 * nlength (4 + 2 * t)" by (simp add: mult.commute) also have "... \ 52 + 1851 * H ^ 4 * (nlength (Suc (Suc (Suc (2 * t)))))\<^sup>2 + 4 * nlllength (nll_Psi (2 * H + H * (2 * t)) H (zs ! t)) + 4 * nlength (2 + 2 * length zs)" using nlength_mono assms(2) by simp also have "... \ 52 + 1851 * H ^ 4 * (nlength (Suc (2 * length zs)))\<^sup>2 + 4 * nlllength (nll_Psi (2 * H + H * (2 * t)) H (zs ! t)) + 4 * nlength (2 + 2 * length zs)" using H4_nlength H assms(2) by simp also have "... \ 52 + 1851 * H ^ 4 * (nlength (Suc (2 * length zs)))\<^sup>2 + 4 * H * (3 + nlength (3 * H + H * (2 * (length zs - 1)))) + 4 * nlength (2 + 2 * length zs)" using nlllength_nll_Psi_le'[of "2 * H + H * (2 * t)" "2 * H + H * (2 * (length zs - 1))" H] * by simp also have "... = 52 + 1851 * H ^ 4 * (nlength (Suc (2 * length zs)))\<^sup>2 + 4 * H * (3 + nlength (H * (1 + 2 * length zs))) + 4 * nlength (2 + 2 * length zs)" proof - have "3 * H + H * (2 * (length zs - 1)) = H * (3 + 2 * (length zs - 1))" by algebra also have "... = H * (3 + 2 * length zs - 2)" using assms(2) by (metis Nat.add_diff_assoc Suc_pred add_mono_thms_linordered_semiring(1) le_numeral_extra(4) length_greater_0_conv list.size(3) mult_2 nat_1_add_1 not_less_zero plus_1_eq_Suc right_diff_distrib' trans_le_add1) also have "... = H * (1 + 2 * length zs)" by simp finally have "3 * H + H * (2 * (length zs - 1)) = H * (1 + 2 * length zs)" . then show ?thesis by metis qed also have "... \ 52 + 1851 * H ^ 4 * (3 + nlength (Suc (2 * length zs)))\<^sup>2 + 4 * H * (3 + nlength (H * (1 + 2 * length zs))) + 4 * nlength (2 + 2 * length zs)" by simp also have "... \ 52 + 1851 * H ^ 4 * (3 + nlength (H * Suc (2 * length zs)))\<^sup>2 + 4 * H * (3 + nlength (H * Suc (2 * length zs))) + 4 * nlength (2 + 2 * length zs)" proof - have "Suc (2 * length zs) \ H * Suc (2 * length zs)" using H_mult by blast then have "nlength (Suc (2 * length zs))^2 \ nlength (H * Suc (2 * length zs))^2" using nlength_mono by simp then show ?thesis by simp qed also have "... \ 52 + 1851 * H ^ 4 * (3 + nlength (H * Suc (2 * length zs)))\<^sup>2 + 4 * H * (3 + nlength (H * Suc (2 * length zs)))^2 + 4 * nlength (2 + 2 * length zs)" using linear_le_pow by simp also have "... \ 52 + 1851 * H ^ 4 * (3 + nlength (H * Suc (2 * length zs)))\<^sup>2 + 4 * H^4 * (3 + nlength (H * Suc (2 * length zs)))^2 + 4 * nlength (2 + 2 * length zs)" using linear_le_pow by simp also have "... = 52 + 1855 * H ^ 4 * (3 + nlength (H * Suc (2 * length zs)))\<^sup>2 + 4 * nlength (2 + 2 * length zs)" by simp also have "... \ 52 + 1855 * H ^ 4 * (3 + nlength (H * Suc (2 * length zs)))\<^sup>2 + 4 * nlength (H * Suc (2 * length zs))" proof - have "2 + 2 * m \ H * Suc (2 * m)" if "m > 0" for m using H H_mult(2) by (metis Suc_leD add_mono eval_nat_numeral(3) mult.commute mult_Suc_right) then have "2 + 2 * length zs \ H * Suc (2 * length zs)" using assms(2) by (metis less_nat_zero_code not_gr_zero) then show ?thesis using nlength_mono by simp qed also have "... \ 52 + 1855 * H ^ 4 * (3 + H * Suc (2 * length zs))\<^sup>2 + 4 * nlength (H * Suc (2 * length zs))" using nlength_le_n nlength_mono by simp also have "... \ 52 + 1855 * H ^ 4 * (3 + H * Suc (2 * length zs))\<^sup>2 + 4 * (H * Suc (2 * length zs))" using nlength_le_n nlength_mono by (meson add_left_mono mult_le_cancel1) also have "... = 52 + 1855 * H ^ 4 * (3^2 + 2 * 3 * H * Suc (2 * length zs) + H^2 * Suc (2 * length zs)^2) + 4 * (H * Suc (2 * length zs))" by algebra also have "... \ 52 + 1855 * H ^ 4 * (72 * H^2 * length zs^2) + 4 * (H * Suc (2 * length zs))" proof - have "3^2 + 2 * 3 * H * Suc (2 * length zs) + H^2 * Suc (2 * length zs)^2 = 9 + 6 * H * Suc (2 * length zs) + H^2 * Suc (2 * length zs)^2" by simp also have "... \ 9 + 6 * H^2 * Suc (2 * length zs) + H^2 * Suc (2 * length zs)^2" using H linear_le_pow by (simp add: add_mono) also have "... \ 9 + 6 * H^2 * Suc (2 * length zs)^2 + H^2 * Suc (2 * length zs)^2" using linear_le_pow by (meson add_le_mono1 add_left_mono le_refl mult_le_mono zero_less_numeral) also have "... = 9 + 7 * H^2 * Suc (2 * length zs)^2" by simp also have "... \ 9 + 7 * H^2 * 9 * length zs^2" using ** by simp also have "... = 9 + 63 * H^2 * length zs^2" by simp also have "... \ 9 * H^2 * length zs^2 + 63 * H^2 * length zs^2" using assms(2) H by simp also have "... = 72 * H^2 * length zs^2" by simp finally have "3^2 + 2 * 3 * H * Suc (2 * length zs) + H^2 * Suc (2 * length zs)^2 \ 72 * H^2 * length zs^2" . then show ?thesis using add_le_mono le_refl mult_le_mono2 by presburger qed also have "... = 52 + 133560 * H ^ 6 * length zs^2 + 4 * (H * Suc (2 * length zs))" by simp also have "... \ 52 + 133560 * H ^ 6 * length zs^2 + 4 * (H * 9 * length zs ^ 2)" using ** by (smt (verit) add_left_mono mult.assoc mult_le_cancel1 power2_nat_le_imp_le) also have "... = 52 + 133560 * H ^ 6 * length zs^2 + 36 * H * length zs ^ 2" by simp also have "... \ 52 + 133560 * H ^ 6 * length zs^2 + 36 * H^6 * length zs ^ 2" using linear_le_pow by simp also have "... = 52 + 133596 * H ^ 6 * length zs^2" by simp also have "... \ 52 * H ^ 6 * length zs^2 + 133596 * H ^ 6 * length zs^2" using H assms(2) by simp also have "... = 133648 * H ^ 6 * length zs^2" by simp finally have "?ttt \ 133648 * H ^ 6 * length zs^2" . then have "transforms tm6 (tpsL t) ttt (tpsL6 t)" using tm6 assms transforms_monotone by blast then show ?thesis using tpsL6_eq_tpsL by simp qed lemma tmL: assumes "ttt = 133650 * H ^ 6 * length zs ^ 3 + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length zs))" unfolding tmL_def proof (tform) let ?t = "133648 * H ^ 6 * length zs^2" show "\i. i < length zs \ transforms tm6 (tpsL i) ?t (tpsL (Suc i))" using tm6' by simp have *: "read (tpsL t) ! 0 = \zs\ (Suc t)" for t using jk tapes_at_read'[symmetric, of 0 "tpsL t"] by (simp add: tpsL_def) show "read (tpsL t) ! 0 \ \" if "t < length zs" for t proof - have "read (tpsL t) ! 0 = zs ! t" using that * by simp then show ?thesis using that zs by auto qed show "\ read (tpsL (length zs)) ! 0 \ \" using * by simp show "length zs * (?t + 2) + 1 \ ttt" proof (cases "length zs = 0") case True then show ?thesis using assms(1) by simp next case False then have "1 \ H^6 * length zs ^ 2" using H linear_le_pow by (simp add: Suc_leI) then have "?t + 2 \ 133650 * H ^ 6 * length zs^2" by simp then have "length zs * (?t + 2) \ length zs * 133650 * H ^ 6 * length zs^2" by simp then have "length zs * (?t + 2) \ 133650 * H ^ 6 * length zs^3" by (simp add: power2_eq_square power3_eq_cube) then show ?thesis using assms(1) by simp qed qed lemma tmL': assumes "ttt = 133650 * H ^ 6 * length zs ^ 3 + 1" shows "transforms tmL tps0 ttt (tpsL (length zs))" using assms tmL tpsL_eq_tps0 by simp end end (* locale turing_machine_PHI6 *) lemma transforms_tm_PHI6I: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k H :: nat and zs :: "symbol list" and nss :: "nat list list" assumes "length tps = k" and "1 < j" and "j + 7 < k" and "H \ 3" and "bit_symbols zs" assumes "tps ! 1 = nlltape nss" "tps ! 0 = (\zs\, 1)" "tps ! j = (\2\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\0\\<^sub>N, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\[]\, 1)" "tps ! (j + 7) = (\[]\, 1)" assumes "tps' = tps [0 := (\zs\, Suc (length zs)), j := (\2 + 2 * length zs\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\i. nll_Psi (H * (2 + 2 * i)) H (zs ! i)) [0..A Turing machine for $\Phi_7$\ text \ The next Turing machine expects a number $idx$ on tape $j$, a number $H$ on tape $j + 1$, and a number $numiter$ on tape $j + 6$. It appends to the CNF formula on tape~$1$ the formula $\bigwedge_{t=0}^{numiter} \Upsilon(\gamma_{idx + 2t})$ with $\gamma_i = [iH, (i+1)H)$. This equals $\Phi_7$ if $idx = 2n + 4$ and $numiter = p(n)$. \ definition tm_PHI7 :: "tapeidx \ machine" where "tm_PHI7 j \ WHILE [] ; \rs. rs ! (j + 6) \ \ DO tm_Upsilongamma j ;; tm_extendl_erase 1 (j + 4) ;; tm_plus_const 2 j ;; tm_decr (j + 6) DONE" lemma tm_PHI7_tm: assumes "0 < j" and "j + 6 < k" and "6 \ G" and "2 \ k" shows "turing_machine k G (tm_PHI7 j)" unfolding tm_PHI7_def using tm_Upsilongamma_tm tm_decr_tm Nil_tm tm_extendl_erase_tm tm_plus_const_tm assms turing_machine_loop_turing_machine by simp locale turing_machine_tm_PHI7 = fixes step :: nat and j :: tapeidx begin definition "tmL1 \ tm_Upsilongamma j" definition "tmL2 \ tmL1 ;; tm_extendl_erase 1 (j + 4)" definition "tmL4 \ tmL2 ;; tm_plus_const 2 j" definition "tmL5 \ tmL4 ;; tm_decr (j + 6)" definition "tmL \ WHILE [] ; \rs. rs ! (j + 6) \ \ DO tmL5 DONE" lemma tmL_eq_tm_PHI7: "tmL = tm_PHI7 j" unfolding tmL_def tmL5_def tmL4_def tmL2_def tmL1_def tm_PHI7_def by simp context fixes tps0 :: "tape list" and numiter H k idx :: nat and nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j" "j + 6 < k" and H: "H \ 3" assumes tps0: "tps0 ! 1 = nlltape nss" "tps0 ! j = (\idx\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\[]\, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\numiter\\<^sub>N, 1)" begin lemma nlength_H: "nlength H \ 1" using nlength_0 H by (metis dual_order.trans nlength_1_simp nlength_mono one_le_numeral) definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j := (\idx + 2 * t\\<^sub>N, 1), j + 6 := (\numiter - t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\t. nll_Upsilon (idx + 2 * t) H) [0.. tape list" where "tpsL1 t \ tps0 [j := (\idx + 2 * t\\<^sub>N, 1), j + 6 := (\numiter - t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\t. nll_Upsilon (idx + 2 * t) H) [0..nll_Upsilon (idx + 2 * t) H\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tmL1 [transforms_intros]: assumes "t < numiter" and "ttt = 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2" shows "transforms tmL1 (tpsL t) ttt (tpsL1 t)" unfolding tmL1_def by (tform tps: H tps0 tpsL_def tpsL1_def jk time: assms(2)) definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [j := (\idx + 2 * t\\<^sub>N, 1), j + 6 := (\numiter - t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\t. nll_Upsilon (idx + 2 * t) H) [0..[]\, 1)]" lemma tmL2 [transforms_intros]: assumes "t < numiter" and "ttt = 11 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * nlllength (nll_Upsilon (idx + 2 * t) H)" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" unfolding tmL2_def by (tform tps: assms(1) H tps0 tpsL1_def tpsL2_def jk time: assms(2)) definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [j := (\idx + 2 * Suc t\\<^sub>N, 1), j + 6 := (\numiter - t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\t. nll_Upsilon (idx + 2 * t) H) [0..[]\, 1)]" lemma tmL4 [transforms_intros]: assumes "t < numiter" and "ttt = 21 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * nlllength (nll_Upsilon (idx + 2 * t) H) + 4 * nlength (Suc (Suc (idx + 2 * t)))" shows "transforms tmL4 (tpsL t) ttt (tpsL4 t)" unfolding tmL4_def proof (tform tps: assms(1) H tps0 tpsL2_def tpsL4_def jk time: assms(2)) show "tpsL4 t = (tpsL2 t)[j := (\idx + 2 * t + 2\\<^sub>N, 1)]" using tpsL4_def tpsL2_def jk by (simp add: list_update_swap[of j]) qed definition tpsL5 :: "nat \ tape list" where "tpsL5 t \ tps0 [j := (\idx + 2 * Suc t\\<^sub>N, 1), j + 6 := (\numiter - Suc t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\t. nll_Upsilon (idx + 2 * t) H) [0..[]\, 1)]" lemma tmL5: assumes "t < numiter" and "ttt = 29 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * nlllength (nll_Upsilon (idx + 2 * t) H) + 4 * nlength (Suc (Suc (idx + 2 * t))) + 2 * nlength (numiter - t)" shows "transforms tmL5 (tpsL t) ttt (tpsL5 t)" unfolding tmL5_def proof (tform tps: assms(1) H tps0 tpsL4_def tpsL5_def jk time: assms(2)) show "tpsL5 t = (tpsL4 t)[j + 6 := (\numiter - t - 1\\<^sub>N, 1)]" using tpsL5_def tpsL4_def jk by (simp add: list_update_swap[of "j+6"]) qed lemma tpsL5_eq_tpsL: "tpsL5 t = tpsL (Suc t)" proof - define tps where "tps = tps0 [j := (\idx + 2 * Suc t\\<^sub>N, 1), j + 6 := (\numiter - Suc t\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\t. nll_Upsilon (idx + 2 * t) H) [0..[]\, 1)" using tps_def jk tps0 by simp then have "tps[j + 4 := (\[]\, 1)] = tps" using list_update_id[of _ "j+4"] by metis then show ?thesis unfolding tps_def using tpsL5_def by simp qed ultimately show ?thesis by simp qed lemma tmL5' [transforms_intros]: assumes "t < numiter" and "ttt = 256 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2" shows "transforms tmL5 (tpsL t) ttt (tpsL (Suc t))" proof - let ?ttt = "29 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * nlllength (nll_Upsilon (idx + 2 * t) H) + 4 * nlength (Suc (Suc (idx + 2 * t))) + 2 * nlength (numiter - t)" have "?ttt \ 29 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * nlllength (nll_Upsilon (idx + 2 * t) H) + 4 * nlength (Suc (Suc (idx + 2 * t))) + 2 * nlength numiter" using nlength_mono by simp also have "... \ 29 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * H * (4 + nlength (idx + 2 * t) + nlength H) + 4 * nlength (Suc (Suc (idx + 2 * t))) + 2 * nlength numiter" using nlllength_nll_Upsilon_le H by simp also have "... \ 29 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * H * (4 + nlength (idx + 2 * numiter) + nlength H) + 4 * nlength (Suc (Suc (idx + 2 * t))) + 2 * nlength numiter" using nlength_mono assms(1) by simp also have "... \ 29 + 205 * H * (nlength (idx + 2 * t) + nlength H)\<^sup>2 + 4 * H * (4 + nlength (idx + 2 * numiter) + nlength H) + 4 * nlength (idx + 2 * numiter) + 2 * nlength numiter" using nlength_mono assms(1) by simp also have "... \ 29 + 205 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 4 * H * (4 + nlength (idx + 2 * numiter) + nlength H) + 4 * nlength (idx + 2 * numiter) + 2 * nlength numiter" using nlength_mono assms(1) by simp also have "... \ 29 + 205 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 4 * H * (4 + nlength (idx + 2 * numiter) + nlength H) + 6 * nlength (idx + 2 * numiter)" using nlength_mono by simp also have "... \ 29 + 205 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 4 * H * (4 + nlength (idx + 2 * numiter) + nlength H) + 6 * (nlength (idx + 2 * numiter) + nlength H)" using nlength_mono by simp also have "... = 29 + 205 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 16 * H + 4 * H * (nlength (idx + 2 * numiter) + nlength H) + 6 * (nlength (idx + 2 * numiter) + nlength H)" by algebra also have "... \ 29 + 205 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 16 * H + 4 * H * (nlength (idx + 2 * numiter) + nlength H) + 2 * H * (nlength (idx + 2 * numiter) + nlength H)" proof - have "6 \ 2 * H" using H by simp then show ?thesis using mult_le_mono1 nat_add_left_cancel_le by presburger qed also have "... = 29 + 205 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 16 * H + 6 * H * (nlength (idx + 2 * numiter) + nlength H)" by simp also have "... \ 29 + 205 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 16 * H + 6 * H * (nlength (idx + 2 * numiter) + nlength H)^2" using linear_le_pow by simp also have "... = 29 + 211 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 16 * H" by simp also have "... \ 29 + 227 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2" using H nlength_0 by (simp add: Suc_leI) also have "... \ 256 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2" using H nlength_0 by (simp add: Suc_leI) finally have "?ttt \ ttt" using assms(2) by simp then have "transforms tmL5 (tpsL t) ttt (tpsL5 t)" using assms(1) tmL5 transforms_monotone by blast then show ?thesis using tpsL5_eq_tpsL by simp qed lemma tmL: assumes "ttt = numiter * (257 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL numiter)" unfolding tmL_def proof (tform) show "\i. i < numiter \ read (tpsL i) ! (j + 6) \ \" using jk tpsL_def read_ncontents_eq_0 by simp show "\ read (tpsL numiter) ! (j + 6) \ \" using jk tpsL_def read_ncontents_eq_0 by simp show "numiter * (256 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 2) + 1 \ ttt" proof - have "1 \ (nlength (idx + 2 * numiter) + nlength H)\<^sup>2" using nlength_H by simp then have "2 \ H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2" using H by (metis add_leE mult_2 mult_le_mono nat_1_add_1 numeral_Bit1 numerals(1)) then show ?thesis using assms by simp qed qed lemma tmL': assumes "ttt = numiter * 257 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 1" shows "transforms tmL tps0 ttt (tpsL numiter)" using assms tmL tpsL_eq_tps0 by (simp add: Groups.mult_ac(1)) end (* context tps0 *) end (* locale turing_machine_tm_PHI7 *) lemma transforms_tm_PHI7I: fixes tps tps' :: "tape list" and ttt numiter H k idx :: nat and nss :: "nat list list" and j :: tapeidx assumes "length tps = k" and "1 < j" and "j + 6 < k" and "H \ 3" assumes "tps ! 1 = nlltape nss" "tps ! j = (\idx\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\[]\, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\numiter\\<^sub>N, 1)" assumes "ttt = numiter * 257 * H * (nlength (idx + 2 * numiter) + nlength H)\<^sup>2 + 1" assumes "tps' = tps [j := (\idx + 2 * numiter\\<^sub>N, 1), j + 6 := (\0\\<^sub>N, 1), 1 := nlltape (nss @ concat (map (\t. nll_Upsilon (idx + 2 * t) H) [0..A Turing machine for $\Phi_8$\ text \ The next TM expects a number $idx$ on tape $j$ and a number $H$ on tape $j + 1$. It appends to the formula on tape~$1$ the formula $\Psi([idx\cdot H, (idx+1)H), 3)$. \ definition tm_PHI8 :: "tapeidx \ machine" where "tm_PHI8 j \ tm_setn (j + 2) 3 ;; tm_Psigamma j ;; tm_extendl 1 (j + 6)" lemma tm_PHI8_tm: assumes "0 < j" and "j + 7 < k" and "G \ 6" shows "turing_machine k G (tm_PHI8 j)" unfolding tm_PHI8_def using assms tm_Psigamma_tm tm_setn_tm tm_extendl_tm by simp locale turing_machine_PHI8 = fixes j :: tapeidx begin definition "tm1 \ tm_setn (j + 2) 3" definition "tm2 \ tm1 ;; tm_Psigamma j" definition "tm3 \ tm2 ;; tm_extendl 1 (j + 6)" lemma tm3_eq_tm_PHI8: "tm3 = tm_PHI8 j" using tm3_def tm2_def tm1_def tm_PHI8_def by simp context fixes tps0 :: "tape list" and k idx H :: nat and nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j" "j + 7 < k" and H: "H \ 3" assumes tps0: "tps0 ! 1 = nlltape nss" "tps0 ! j = (\idx\\<^sub>N, 1)" "tps0 ! (j + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j + 2) = (\[]\, 1)" "tps0 ! (j + 3) = (\[]\, 1)" "tps0 ! (j + 4) = (\[]\, 1)" "tps0 ! (j + 5) = (\[]\, 1)" "tps0 ! (j + 6) = (\[]\, 1)" "tps0 ! (j + 7) = (\[]\, 1)" begin definition "tps1 \ tps0 [j + 2 := (\3\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 14" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps0 tps1_def jk) show "tps0 ! (j + 2) = (\0\\<^sub>N, 1)" using tps0 jk canrepr_0 by simp show "ttt = 10 + 2 * nlength 0 + 2 * nlength 3" using assms nlength_3 by simp qed definition "tps2 \ tps0 [j + 2 := (\3\\<^sub>N, 1), j + 6 := (\nll_Psi (idx * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 14 + 1851 * H ^ 4 * nlength (Suc idx) ^ 2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: tps0 H tps1_def tps2_def jk time: assms canrepr_1) definition "tps3 \ tps0 [1 := nlltape (nss @ nll_Psi (idx * H) H 3), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\nll_Psi (idx * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1)]" lemma tm3: assumes "ttt = 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * nlllength (nll_Psi (idx * H) H 3)" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def by (tform tps: tps0 H tps2_def tps3_def jk time: assms) lemma tm3': assumes "ttt = 18 + 1861 * H ^ 4 * (nlength (Suc idx))\<^sup>2" shows "transforms tm3 tps0 ttt tps3" proof - have *: "(nlength (Suc idx))\<^sup>2 \ 1" using nlength_0 by (simp add: Suc_leI) let ?ttt = "18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * nlllength (nll_Psi (idx * H) H 3)" have "?ttt \ 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * H * (3 + nlength (idx * H + H))" using nlllength_nll_Psi_le by simp also have "... = 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * H * (3 + nlength (Suc idx * H))" by (simp add: add.commute) also have "... \ 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * H * (3 + nlength (Suc idx) + nlength H)" using nlength_prod by (metis (mono_tags, lifting) ab_semigroup_add_class.add_ac(1) add_left_mono mult_le_cancel1) also have "... = 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 6 * H + 2 * H * nlength (Suc idx) + 2 * H * nlength H" by algebra also have "... \ 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 6 * H + 2 * H * nlength (Suc idx) + 2 * H * H" using nlength_le_n by simp also have "... \ 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 6 * H ^ 4 + 2 * H * nlength (Suc idx) + 2 * H * H" using linear_le_pow[of 4 H] by simp also have "... \ 18 + 1851 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 6 * H ^ 4 + 2 * H ^ 4 * nlength (Suc idx) + 2 * H * H" using linear_le_pow[of 4 H] by simp also have "... \ 18 + 1857 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * H ^ 4 * nlength (Suc idx) + 2 * H * H" using * by simp also have "... \ 18 + 1859 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * H * H" using linear_le_pow[of 2 "nlength (Suc idx)"] by simp also have "... = 18 + 1859 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * H ^ 2" by algebra also have "... \ 18 + 1859 * H ^ 4 * (nlength (Suc idx))\<^sup>2 + 2 * H ^ 4" using pow_mono'[of 2 4 H] by simp also have "... \ 18 + 1861 * H ^ 4 * (nlength (Suc idx))\<^sup>2" using * by simp finally have "?ttt \ 18 + 1861 * H ^ 4 * (nlength (Suc idx))\<^sup>2" . then show ?thesis using tm3 assms transforms_monotone by simp qed end end (* locale turing_machine_PHI8 *) lemma transforms_tm_PHI8I: fixes j :: tapeidx fixes tps tps' :: "tape list" and ttt k idx H :: nat and nss :: "nat list list" assumes "length tps = k" and "1 < j" and "j + 7 < k" and "H \ 3" assumes "tps ! 1 = nlltape nss" "tps ! j = (\idx\\<^sub>N, 1)" "tps ! (j + 1) = (\H\\<^sub>N, 1)" "tps ! (j + 2) = (\[]\, 1)" "tps ! (j + 3) = (\[]\, 1)" "tps ! (j + 4) = (\[]\, 1)" "tps ! (j + 5) = (\[]\, 1)" "tps ! (j + 6) = (\[]\, 1)" "tps ! (j + 7) = (\[]\, 1)" assumes "tps' = tps [1 := nlltape (nss @ nll_Psi (idx * H) H 3), j + 2 := (\3\\<^sub>N, 1), j + 6 := (\nll_Psi (idx * H) H 3\\<^sub>N\<^sub>L\<^sub>L, 1)]" assumes "ttt = 18 + 1861 * H ^ 4 * (nlength (Suc idx))\<^sup>2" shows "transforms (tm_PHI8 j) tps ttt tps'" proof - interpret loc: turing_machine_PHI8 j . show ?thesis using loc.tps3_def loc.tm3' loc.tm3_eq_tm_PHI8 assms by metis qed subsection \A Turing machine for $\Phi_9$\ text \ The CNF formula $\Phi_9 = \bigwedge_{t=1}^{T'}$ is the most complicated part of $\Phi$. Clearly, the main task here is to generate the formulas $\chi_t$ \ subsubsection \A Turing machine for $\chi_t$\ text \ A lemma that will help with some time bounds: \ lemma pow2_le_2pow2: "z ^ 2 \ 2 ^ (2*z)" for z :: nat proof (induction z) case 0 then show ?case by simp next case (Suc z) show ?case proof (cases "z = 0") case True then show ?thesis by simp next case False have "Suc z ^ 2 = z ^ 2 + 2 * z + 1" by (metis Suc_eq_plus1 ab_semigroup_add_class.add_ac(1) nat_mult_1_right one_power2 plus_1_eq_Suc power2_sum) also have "... \ z ^ 2 + 3 * z" using False by simp also have "... \ z ^ 2 + 3 * z ^ 2" by (simp add: linear_le_pow) also have "... = 2^2 * z ^ 2" by simp also have "... \ 2^2 * 2 ^ (2 * z)" using Suc by simp also have "... = 2 ^ (2 * Suc z)" by (simp add: power_add) finally show ?thesis . qed qed text \ The next Turing machine can be used to generate $\chi_t$. It expects on tape~1 a CNF formula, on tape~$j_1$ the list of positions of $M$'s input tape head, on tape~$j_2$ the list of positions of $M$'s output tape head, on tapes~$j_3, \dots, j_3+3$ the numbers $N$, $G$, $Z$, and $T$, on tape~$j_3+4$ the formula $\psi$, on tape~$j_3+5$ the formula $\psi'$, and finally on tape~$j_3+6$ the number $t$. The TM appends the formula $\chi_t$ to the formula on tape~1, which should be thought of as an unfinished version of $\Phi$. The TM first computes $\prev(t)$ using @{const tm_prev} and compares it with $t$. Depending on the outcome of this comparison it generates either $\rho_t$ or $\rho'_t$ by concatenating ranges of numbers generated using @{const tm_range}. Then the TM uses @{const tm_relabel} to compute $\rho_t(\psi)$ or $\rho'_t(\psi')$. The result equals $\chi_t$ and is appended to tape~1. Finally $t$ is incremented and $T$ is decremented. This is so the TM can be used inside a while loop that initializes $T$ with $T'$ and $t$ with $1$. \null \ definition tm_chi :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_chi j1 j2 j3 \ tm_prev j2 (j3 + 6) ;; tm_equalsn (j3 + 11) (j3 + 6) (j3 + 13) ;; tm_decr (j3 + 6) ;; tm_mult (j3 + 6) (j3 + 2) (j3 + 7) ;; tm_add j3 (j3 + 7) ;; tm_range (j3 + 7) (j3 + 2) (j3 + 8) ;; tm_extend_erase (j3 + 12) (j3 + 8) ;; tm_setn (j3 + 7) 0 ;; IF \rs. rs ! (j3 + 13) = \ THEN tm_mult (j3 + 11) (j3 + 2) (j3 + 7) ;; tm_add j3 (j3 + 7) ;; tm_range (j3 + 7) (j3 + 2) (j3 + 8) ;; tm_extend_erase (j3 + 12) (j3 + 8) ;; tm_setn (j3 + 7) 0 ELSE [] ENDIF ;; tm_incr (j3 + 6) ;; tm_mult (j3 + 6) (j3 + 2) (j3 + 7) ;; tm_add j3 (j3 + 7) ;; tm_range (j3 + 7) (j3 + 2) (j3 + 8) ;; tm_extend_erase (j3 + 12) (j3 + 8) ;; tm_setn (j3 + 11) 0 ;; tm_nth j1 (j3 + 6) (j3 + 11) 4 ;; tm_setn (j3 + 7) 0 ;; tm_mult (j3 + 11) (j3 + 1) (j3 + 7) ;; tm_range (j3 + 7) (j3 + 1) (j3 + 8) ;; tm_extend_erase (j3 + 12) (j3 + 8) ;; tm_setn (j3 + 7) 0 ;; tm_erase_cr (j3 + 11) ;; tm_cr (j3 + 12) ;; IF \rs. rs ! (j3 + 13) = \ THEN tm_relabel (j3 + 4) (j3 + 11) ELSE tm_erase_cr (j3 + 13) ;; tm_relabel (j3 + 5) (j3 + 11) ENDIF ;; tm_erase_cr (j3 + 12) ;; tm_extendl_erase 1 (j3 + 11) ;; tm_incr (j3 + 6) ;; tm_decr (j3 + 3)" lemma tm_chi_tm: assumes "0 < j1" and "j1 < j2" and "j2 < j3" and "j3 + 17 < k" and "G \ 6" shows "turing_machine k G (tm_chi j1 j2 j3)" unfolding tm_chi_def using tm_prev_tm tm_equalsn_tm tm_decr_tm tm_mult_tm tm_add_tm tm_range_tm tm_extend_erase_tm tm_setn_tm tm_incr_tm tm_nth_tm tm_erase_cr_tm tm_relabel_tm Nil_tm tm_cr_tm tm_extendl_erase_tm assms turing_machine_branch_turing_machine by simp locale turing_machine_chi = fixes j1 j2 j3 :: tapeidx begin definition "tm1 \ tm_prev j2 (j3 + 6)" definition "tm2 \ tm1 ;; tm_equalsn (j3 + 11) (j3 + 6) (j3 + 13)" definition "tm3 \ tm2 ;; tm_decr (j3 + 6)" definition "tm4 \ tm3 ;; tm_mult (j3 + 6) (j3 + 2) (j3 + 7)" definition "tm5 \ tm4 ;; tm_add j3 (j3 + 7)" definition "tm6 \ tm5 ;; tm_range (j3 + 7) (j3 + 2) (j3 + 8)" definition "tm7 \ tm6 ;; tm_extend_erase (j3 + 12) (j3 + 8)" definition "tm8 \ tm7 ;; tm_setn (j3 + 7) 0" definition "tmT1 \ tm_mult (j3 + 11) (j3 + 2) (j3 + 7)" definition "tmT2 \ tmT1 ;; tm_add j3 (j3 + 7)" definition "tmT3 \ tmT2 ;; tm_range (j3 + 7) (j3 + 2) (j3 + 8)" definition "tmT4 \ tmT3 ;; tm_extend_erase (j3 + 12) (j3 + 8)" definition "tmT5 \ tmT4 ;; tm_setn (j3 + 7) 0" definition "tm89 \ IF \rs. rs ! (j3 + 13) = \ THEN tmT5 ELSE [] ENDIF" definition "tm10 \ tm8 ;; tm89" definition "tm11 \ tm10 ;; tm_incr (j3 + 6)" definition "tm13 \ tm11 ;; tm_mult (j3 + 6) (j3 + 2) (j3 + 7)" definition "tm14 \ tm13 ;; tm_add j3 (j3 + 7)" definition "tm15 \ tm14 ;; tm_range (j3 + 7) (j3 + 2) (j3 + 8)" definition "tm16 \ tm15 ;; tm_extend_erase (j3 + 12) (j3 + 8)" definition "tm17 \ tm16 ;; tm_setn (j3 + 11) 0" definition "tm18 \ tm17 ;; tm_nth j1 (j3 + 6) (j3 + 11) 4" definition "tm19 \ tm18 ;; tm_setn (j3 + 7) 0" definition "tm20 \ tm19 ;; tm_mult (j3 + 11) (j3 + 1) (j3 + 7)" definition "tm21 \ tm20 ;; tm_range (j3 + 7) (j3 + 1) (j3 + 8)" definition "tm22 \ tm21 ;; tm_extend_erase (j3 + 12) (j3 + 8)" definition "tm23 \ tm22 ;; tm_setn (j3 + 7) 0" definition "tm24 \ tm23 ;; tm_erase_cr (j3 + 11)" definition "tm25 \ tm24 ;; tm_cr (j3 + 12)" definition "tmE1 \ tm_erase_cr (j3 + 13)" definition "tmE2 \ tmE1 ;; tm_relabel (j3 + 5) (j3 + 11)" definition "tmTT1 \ tm_relabel (j3 + 4) (j3 + 11)" definition "tm2526 \ IF \rs. rs ! (j3 + 13) = \ THEN tmTT1 ELSE tmE2 ENDIF" definition "tm26 \ tm25 ;; tm2526" definition "tm27 \ tm26 ;; tm_erase_cr (j3 + 12)" definition "tm28 \ tm27 ;; tm_extendl_erase 1 (j3 + 11)" definition "tm29 \ tm28 ;; tm_incr (j3 + 6)" definition "tm30 \ tm29 ;; tm_decr (j3 + 3)" lemma tm30_eq_tm_chi: "tm30 = tm_chi j1 j2 j3" unfolding tm30_def tm29_def tm28_def tm27_def tm26_def tm25_def tm24_def tm23_def tm22_def tm21_def tm20_def tm19_def tm18_def tm17_def tm16_def tm15_def tm14_def tm13_def tm11_def tm10_def tm8_def tm7_def tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def tm89_def tmE1_def tmE2_def tmTT1_def tm2526_def tmT5_def tmT4_def tmT3_def tmT2_def tmT1_def tm_chi_def by simp context fixes tps0 :: "tape list" and k G N Z T' T t :: nat and hp0 hp1 :: "nat list" and \ \' :: formula fixes nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j1" "j1 < j2" "j2 < j3" "j3 + 17 < k" and G: "G \ 3" and Z: "Z = 3 * G" and N: "N \ 1" and len_hp0: "length hp0 = Suc T'" and hp0: "\i T'" and len_hp1: "length hp1 = Suc T'" and hp1: "\i T'" and t: "0 < t" "t \ T'" and T: "0 < T" "T \ T'" and T': "T' < N" and psi: "variables \ \ {..<3*Z+G}" "fsize \ \ (3*Z+G) * 2 ^ (3*Z+G)" "length \ \ 2 ^ (3*Z+G)" and psi': "variables \' \ {..<2*Z+G}" "fsize \' \ (2*Z+G) * 2 ^ (2*Z+G)" "length \' \ 2 ^ (2*Z+G)" assumes tps0: "tps0 ! 1 = nlltape nss" "tps0 ! j1 = (\hp0\\<^sub>N\<^sub>L, 1)" "tps0 ! j2 = (\hp1\\<^sub>N\<^sub>L, 1)" "tps0 ! j3 = (\N\\<^sub>N, 1)" "tps0 ! (j3 + 1) = (\G\\<^sub>N, 1)" "tps0 ! (j3 + 2) = (\Z\\<^sub>N, 1)" "tps0 ! (j3 + 3) = (\T\\<^sub>N, 1)" "tps0 ! (j3 + 4) = (\formula_n \\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j3 + 5) = (\formula_n \'\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j3 + 6) = (\t\\<^sub>N, 1)" "\i. 6 < i \ i < 17 \ tps0 ! (j3 + i) = (\[]\, 1)" begin lemma Z_ge_1: "Z \ 1" using G Z by simp lemma Z_ge_9: "Z \ 9" using G Z by simp lemma T'_ge_1: "T' \ 1" using t by simp lemma tps0': "\i. 1 \ i \ i < 11 \ tps0 ! (j3 + 6 + i) = (\[]\, 1)" proof - fix i :: nat assume i: "1 \ i" "i < 11" then have "6 < 6 + i" "6 + i < 17" by simp_all then have "tps0 ! (j3 + (6 + i)) = (\[]\, 1)" using tps0(11) by simp then show "tps0 ! (j3 + 6 + i) = (\[]\, 1)" by (metis group_cancel.add1) qed text \The simplifier turns $j3 + 6 + 3$ into $9 + j3$. The next lemma helps with that.\ lemma tps0_sym: "\i. 6 < i \ i < 17 \ tps0 ! (i + j3) = (\[]\, 1)" using tps0(11) by (simp add: add.commute) lemma previous_hp1_le: "previous hp1 t \ t" using len_hp1 hp1 t(2) previous_le[of hp1 t] by simp definition "tps1 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 71 + 153 * nllength hp1 ^ 3" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: canrepr_0 tps0_sym tps0 tps1_def jk t len_hp1 time: assms) show "tps1 = tps0[j3 + 6 + 5 := (\previous hp1 t\\<^sub>N, 1)]" using tps1_def by (simp add: add.commute) qed definition "tps2 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 78 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t)" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps0 tps1_def tps2_def jk time: assms) show "tps1 ! (j3 + 13) = (\0\\<^sub>N, 1)" using tps1_def tps0(11) canrepr_0 by simp show "(0::nat) \ 1" by simp qed definition "tps3 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 86 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength t" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def by (tform tps: tps0 tps2_def tps3_def jk assms) definition "tps4 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\(t - 1) * Z\\<^sub>N, 1)]" lemma tm4 [transforms_intros]: assumes "ttt = 90 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength t + 26 * (nlength (t - 1) + nlength Z) ^ 2" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: tps0 tps3_def tps4_def jk) show "tps3 ! (j3 + 7) = (\0\\<^sub>N, 1)" using tps3_def jk canrepr_0 tps0 by simp show "ttt = 86 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength t + (4 + 26 * (nlength (t - Suc 0) + nlength Z) * (nlength (t - Suc 0) + nlength Z))" proof - have "(26 * nlength (t - Suc 0) + 26 * nlength Z) * (nlength (t - Suc 0) + nlength Z) = 26 * (nlength (t - Suc 0) + nlength Z) ^ 2" by algebra then show ?thesis using assms by simp qed qed definition "tps5 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\N + (t - 1) * Z\\<^sub>N, 1)]" lemma tm5 [transforms_intros]: assumes "ttt = 100 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength t + 26 * (nlength (t - 1) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z))" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def by (tform tps: tps0 tps4_def tps5_def jk time: assms) definition "tps6 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\N + (t - 1) * Z\\<^sub>N, 1), j3 + 8 := (\[N + (t - 1) * Z..\<^sub>N\<^sub>L, 1)]" lemma tm6 [transforms_intros]: assumes "ttt = 100 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength t + 26 * (nlength (t - 1) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z)) + Suc Z * (43 + 9 * nlength (N + (t - 1) * Z + Z))" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def by (tform tps: nlcontents_Nil canrepr_0 tps0_sym tps0 tps5_def tps6_def jk time: assms) definition "tps7 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\N + (t - 1) * Z\\<^sub>N, 1), j3 + 12 := nltape [N + (t - 1) * Z..[]\, 1)]" unfolding tps7_def tps6_def using tps0 jk list_update_id[of tps0 "j3 + 8"] by (simp add: list_update_swap) qed definition "tps8 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape [N + (t - 1) * Z.. For the next upper bound we have no scruples replacing $\log T'$, $\log N$, and $\log Z$ by $T'$, $N$ and $Z$, respectively. All values are polynomial in $n$ ($Z$ is even a constant), so the overall polynomiality is not in jeopardy. \ lemma nllength_le: fixes nmax :: nat and ns :: "nat list" assumes "\n\set ns. n \ nmax" shows "nllength ns \ Suc nmax * length ns" proof - have "nllength ns \ Suc (nlength nmax) * length ns" using assms nllength_le_len_mult_max by simp then show ?thesis using nlength_le_n by (meson Suc_le_mono dual_order.trans mult_le_mono1) qed lemma nllength_upt_le: fixes a b :: nat shows "nllength [a.. Suc b * (b - a)" proof - have "nllength [a.. Suc (nlength b) * (b - a)" using nllength_upt_le_len_mult_max by simp then show ?thesis using nlength_le_n by (meson Suc_le_mono dual_order.trans mult_le_mono1) qed lemma nllength_hp1: "nllength hp1 \ Suc T' * Suc T'" proof - have "\n\set hp1. n \ T'" using hp1 by (metis in_set_conv_nth) then have "nllength hp1 \ Suc T' * Suc T'" using nllength_le[of hp1 T'] len_hp1 by simp then show ?thesis by simp qed definition "ttt8 \ 168 + 153 * Suc T' ^ 6 + 5 * t + 26 * (t + Z) ^ 2 + 47 * Z + 15 * Z * (N + t * Z)" lemma tm8' [transforms_intros]: "transforms tm8 tps0 ttt8 tps8" proof - let ?s = "88 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength (N + (t - 1) * Z)" let ?ttt = "121 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength t + 26 * (nlength (t - 1) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z)) + Suc Z * (43 + 9 * nlength (N + (t - 1) * Z + Z)) + 4 * nllength [N + (t - 1) * Z.. ?s + 33 + 2 * t + 26 * (nlength (t - 1) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z)) + Suc Z * (43 + 9 * nlength (N + (t - 1) * Z + Z)) + 4 * nllength [N + (t - 1) * Z.. ?s + 33 + 2 * t + 26 * ((t - 1) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z)) + Suc Z * (43 + 9 * nlength (N + (t - 1) * Z + Z)) + 4 * nllength [N + (t - 1) * Z.. ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z)) + Suc Z * (43 + 9 * nlength (N + (t - 1) * Z + Z)) + 4 * nllength [N + (t - 1) * Z.. ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z)) + Suc Z * (43 + 9 * (N + (t - 1) * Z + Z)) + 4 * nllength [N + (t - 1) * Z.. ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * max (nlength N) (nlength ((t - 1) * Z)) + Suc Z * (43 + 9 * (N + (t - 1) * Z + Z)) + 4 * Suc (nlength (N + (t - 1) * Z + Z)) * Z" proof - have "nllength [N + (t - 1) * Z.. Suc (nlength (N + (t - 1) * Z + Z)) * Z" using nllength_le_len_mult_max[of "[N + (t - 1) * Z.. ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * max N ((t - 1) * Z) + Suc Z * (43 + 9 * (N + (t - 1) * Z + Z)) + 4 * Suc (nlength (N + (t - 1) * Z + Z)) * Z" using nlength_le_n by simp also have "... = ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * max N ((t - 1) * Z) + Suc Z * (43 + 9 * (N + t * Z)) + 4 * Suc (nlength (N + (t - 1) * Z + Z)) * Z" using t by (smt (verit) ab_semigroup_add_class.add_ac(1) add.commute diff_Suc_1 gr0_implies_Suc mult_Suc) also have "... \ ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * max N ((t - 1) * Z) + Suc Z * (43 + 9 * (N + t * Z)) + 4 * Suc (N + (t - 1) * Z + Z) * Z" using nlength_le_n Suc_le_mono add_le_mono le_refl mult_le_mono by presburger also have "... = ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * max N ((t - 1) * Z) + Suc Z * (43 + 9 * (N + t * Z)) + 4 * Suc (N + t * Z) * Z" by (smt (verit, del_insts) Suc_eq_plus1 Suc_leI add.commute add.left_commute add_leD2 le_add_diff_inverse mult.commute mult_Suc_right t(1)) also have "... \ ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + ((t - 1) * Z)) + Suc Z * (43 + 9 * (N + t * Z)) + 4 * Suc (N + t * Z) * Z" by simp also have "... \ ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (43 + 9 * (N + t * Z)) + 4 * Suc (N + t * Z) * Z" by simp also have "... = ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (43 + 9 * (N + t * Z)) + (4 + 4 * (N + t * Z)) * Z" by simp also have "... \ ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (43 + 9 * (N + t * Z)) + (4 + 4 * (N + t * Z)) * Suc Z" by simp also have "... = ?s + 33 + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" by algebra also have "... = 121 + 153 * nllength hp1 ^ 3 + 3 * nlength (min (previous hp1 t) t) + 2 * nlength (N + (t - 1) * Z) + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" by simp also have "... \ 121 + 153 * nllength hp1 ^ 3 + 3 * nlength t + 2 * nlength (N + (t - 1) * Z) + 2 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" using previous_hp1_le nlength_mono by simp also have "... \ 121 + 153 * nllength hp1 ^ 3 + 2 * nlength (N + (t - 1) * Z) + 5 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" using nlength_le_n by simp also have "... \ 121 + 153 * nllength hp1 ^ 3 + 2 * (N + (t - 1) * Z) + 5 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" using nlength_le_n add_le_mono1 mult_le_mono2 nat_add_left_cancel_le by presburger also have "... \ 121 + 153 * nllength hp1 ^ 3 + 2 * (N + t * Z) + 5 * t + 26 * ((t - 1) + Z) ^ 2 + 3 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" by simp also have "... = 121 + 153 * nllength hp1 ^ 3 + 5 * t + 26 * ((t - 1) + Z) ^ 2 + 5 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" by simp also have "... \ 121 + 153 * (Suc T' * Suc T') ^ 3 + 5 * t + 26 * ((t - 1) + Z) ^ 2 + 5 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" using nllength_hp1 by simp also have "... = 121 + 153 * (Suc T' ^ 2)^3 + 5 * t + 26 * ((t - 1) + Z) ^ 2 + 5 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" by algebra also have "... = 121 + 153 * Suc T' ^ 6 + 5 * t + 26 * ((t - 1) + Z) ^ 2 + 5 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" by simp also have "... \ 121 + 153 * Suc T' ^ 6 + 5 * t + 26 * (t + Z) ^ 2 + 5 * (N + t * Z) + Suc Z * (47 + 13 * (N + t * Z))" by simp also have "... = 121 + 153 * Suc T' ^ 6 + 5 * t + 26 * (t + Z) ^ 2 + 5 * (N + t * Z) + 47 + 13 * (N + t * Z) + Z * (47 + 13 * (N + t * Z))" by simp also have "... = 168 + 153 * Suc T' ^ 6 + 5 * t + 26 * (t + Z) ^ 2 + 18 * (N + t * Z) + Z * (47 + 13 * (N + t * Z))" by simp also have "... \ 168 + 153 * Suc T' ^ 6 + 5 * t + 26 * (t + Z) ^ 2 + 2 * Z * (N + t * Z) + Z * (47 + 13 * (N + t * Z))" using Z_ge_9 by (metis add_le_mono add_le_mono1 mult_2 mult_le_mono1 nat_add_left_cancel_le numeral_Bit0) also have "... = 168 + 153 * Suc T' ^ 6 + 5 * t + 26 * (t + Z) ^ 2 + 2 * Z * (N + t * Z) + 47 * Z + 13 * Z * (N + t * Z)" by algebra also have "... = 168 + 153 * Suc T' ^ 6 + 5 * t + 26 * (t + Z) ^ 2 + 47 * Z + 15 * Z * (N + t * Z)" by simp finally have "?ttt \ ttt8" using ttt8_def by simp then show ?thesis using tm8 transforms_monotone by simp qed definition "tpsT1 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\previous hp1 t * Z\\<^sub>N, 1), j3 + 12 := nltape [N + (t - 1) * Z..previous hp1 t * Z\\<^sub>N, 1)]" unfolding tpsT1_def tps8_def by (simp add: list_update_swap[of "j3+7"]) qed definition "tpsT2 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\N + previous hp1 t * Z\\<^sub>N, 1), j3 + 12 := nltape [N + (t - 1) * Z.. tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\N + previous hp1 t * Z\\<^sub>N, 1), j3 + 12 := nltape [N + (t - 1) * Z..[N + previous hp1 t * Z..\<^sub>N\<^sub>L, 1)]" lemma tmT3 [transforms_intros]: assumes "ttt = 14 + 26 * (nlength (previous hp1 t) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z))" shows "transforms tmT3 tps8 ttt tpsT3" unfolding tmT3_def by (tform tps: nlcontents_Nil canrepr_0 tps0_sym tps0 tpsT2_def tpsT3_def jk time: assms) definition "tpsT4 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\N + previous hp1 t * Z\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - 1) * Z..[]\, 1)]" lemma tmT4 [transforms_intros]: assumes "ttt = 25 + 26 * (nlength (previous hp1 t) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 4 * nllength [N + previous hp1 t * Z.. tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - 1) * Z..[]\, 1)]" lemma tmT5 [transforms_intros]: assumes "ttt = 35 + 26 * (nlength (previous hp1 t) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 4 * nllength [N + previous hp1 t * Z.. tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - 1) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm89 [transforms_intros]: assumes "ttt = 37 + 26 * (nlength (previous hp1 t) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 4 * nllength [N + previous hp1 t * Z..previous hp1 t = t\\<^sub>B, 1)" using tps8_def jk by simp then have *: "(previous hp1 t \ t) = (read tps8 ! (j3 + 13) = \)" using jk tps8_def read_ncontents_eq_0 by force show "read tps8 ! (j3 + 13) = \ \ tps9 = tpsT5" using * tps9_def tpsT5_def by simp have "(\[]\, 1) = tps0 ! (j3 + 8)" using jk tps0 by simp then have "tps0[j3 + 8 := (\[]\, 1)] = tps0" using jk tps0 by (metis list_update_id) then have "tps8 = tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t - 1\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape [N + (t - 1) * Z..[]\, 1)]" using tps8_def jk tps0 by (simp add: list_update_swap[of _ "j3 + 8"]) then show "read tps8 ! (j3 + 13) \ \ \ tps9 = tps8" using * tps9_def by simp qed definition "ttt10 \ ttt8 + 37 + 26 * (nlength (previous hp1 t) + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 4 * nllength [N + previous hp1 t * Z.. tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - 1) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm11 [transforms_intros]: assumes "ttt = ttt10 + 5 + 2 * nlength (t - 1)" shows "transforms tm11 tps0 ttt tps11" unfolding tm11_def by (tform tps: t(1) tps0 tps9_def tps11_def jk time: assms) definition "tps13 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\t * Z\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - 1) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm13 [transforms_intros]: assumes "ttt = ttt10 + 9 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2" shows "transforms tm13 tps0 ttt tps13" unfolding tm13_def proof (tform tps: tps0 tps11_def tps13_def jk) show "ttt = ttt10 + 5 + 2 * nlength (t - 1) + (4 + 26 * (nlength t + nlength Z) * (nlength t + nlength Z))" using assms by simp (metis Groups.mult_ac(1) distrib_left power2_eq_square) show "tps13 = tps11[j3 + 7 := (\t * Z\\<^sub>N, 1)]" unfolding tps13_def tps11_def by (simp add: list_update_swap[of "j3+7"]) qed definition "tps14 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\N + t * Z\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - 1) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm14 [transforms_intros]: assumes "ttt = ttt10 + 19 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (t * Z))" shows "transforms tm14 tps0 ttt tps14" unfolding tm14_def proof (tform tps: tps0 tps13_def tps14_def jk time: assms) show "tps14 = tps13[j3 + 7 := (\N + t * Z\\<^sub>N, 1)]" unfolding tps14_def tps13_def by (simp add: list_update_swap[of "j3+7"]) qed definition "tps15 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\N + t * Z\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - 1) * Z.. t then [N + previous hp1 t * Z..[N + t * Z..\<^sub>N\<^sub>L, 1)]" lemma tm15 [transforms_intros]: assumes "ttt = ttt10 + 19 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z))" shows "transforms tm15 tps0 ttt tps15" unfolding tm15_def proof (tform tps: tps0 tps14_def tps15_def jk time: assms) show "tps14 ! (j3 + 8) = (\[]\\<^sub>N\<^sub>L, 1)" using tps14_def jk nlcontents_Nil tps0 by simp show "tps14 ! (j3 + 8 + 1) = (\0\\<^sub>N, 1)" using tps14_def jk canrepr_0 tps0_sym by simp show "tps14 ! (j3 + 8 + 2) = (\0\\<^sub>N, 1)" using tps14_def jk canrepr_0 tps0_sym by simp qed definition "tps16 \ tps0 [j3 + 11 := (\previous hp1 t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\N + t * Z\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - Suc 0) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm16 [transforms_intros]: assumes "ttt = ttt10 + 30 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z.. tps0 [j3 + 11 := (\0\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\N + t * Z\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - Suc 0) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm17 [transforms_intros]: assumes "ttt = ttt10 + 40 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z.. tps0 [j3 + 11 := (\hp0 ! t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\N + t * Z\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - Suc 0) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm18 [transforms_intros]: assumes "ttt = ttt10 + 66 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z..2" shows "transforms tm18 tps0 ttt tps18" unfolding tm18_def proof (tform tps: tps0 tps18_def tps17_def jk time: assms) show "t < length hp0" using len_hp0 t by simp qed definition "tps19 \ tps0 [j3 + 11 := (\hp0 ! t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - Suc 0) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" lemma tm19 [transforms_intros]: assumes "ttt = ttt10 + 76 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z)" shows "transforms tm19 tps0 ttt tps19" unfolding tm19_def by (tform tps: tps0 tps18_def tps19_def jk time: assms) definition "tps20 \ tps0 [j3 + 11 := (\hp0 ! t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\hp0 ! t * G\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - Suc 0) * Z.. t then [N + previous hp1 t * Z..[]\, 1)]" definition "ttt20 \ ttt10 + 80 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z) ^ 2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z) + 26 * (nlength (hp0 ! t) + nlength G) ^ 2" lemma tm20 [transforms_intros]: "transforms tm20 tps0 ttt20 tps20" unfolding tm20_def proof (tform tps: tps0 tps19_def tps20_def jk) show "tps20 = tps19[j3 + 7 := (\hp0 ! t * G\\<^sub>N, 1)]" unfolding tps20_def tps19_def by (simp add: list_update_swap[of "j3 + 7"]) show "ttt20 = ttt10 + 76 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z)\<^sup>2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z) + (4 + 26 * (nlength (hp0 ! t) + nlength G) * (nlength (hp0 ! t) + nlength G))" using ttt20_def by simp (metis Groups.mult_ac(1) distrib_left power2_eq_square) qed definition "tps21 \ tps0 [j3 + 11 := (\hp0 ! t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\hp0 ! t * G\\<^sub>N, 1), j3 + 12 := nltape ([N + (t - Suc 0) * Z.. t then [N + previous hp1 t * Z..[hp0 ! t * G..\<^sub>N\<^sub>L, 1)]" lemma tm21 [transforms_intros]: assumes "ttt = ttt20 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G))" shows "transforms tm21 tps0 ttt tps21" unfolding tm21_def proof (tform tps: tps0 tps20_def tps21_def jk time: assms) show "tps20 ! (j3 + 8) = (\[]\\<^sub>N\<^sub>L, 1)" using tps20_def jk nlcontents_Nil tps0 by simp show "tps20 ! (j3 + 8 + 1) = (\0\\<^sub>N, 1)" using tps20_def jk canrepr_0 tps0_sym by simp show "tps20 ! (j3 + 8 + 2) = (\0\\<^sub>N, 1)" using tps20_def jk canrepr_0 tps0_sym by simp qed abbreviation "\ \ [N + (t - 1) * Z.. t then [N + previous hp1 t * Z.. tps0 [j3 + 11 := (\hp0 ! t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\hp0 ! t * G\\<^sub>N, 1), j3 + 12 := nltape \, j3 + 8 := (\[]\, 1)]" lemma tm22 [transforms_intros]: assumes "ttt = ttt20 + 11 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. tps0 [j3 + 11 := (\hp0 ! t\\<^sub>N, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape \, j3 + 8 := (\[]\, 1)]" lemma tm23 [transforms_intros]: assumes "ttt = ttt20 + 21 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. tps0 [j3 + 11 := (\[]\, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := nltape \, j3 + 8 := (\[]\, 1)]" lemma tm24 [transforms_intros]: assumes "ttt = ttt20 + 28 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. tps0 [j3 + 11 := (\[]\, 1), j3 + 13 := (\previous hp1 t = t\\<^sub>B, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := (\\\\<^sub>N\<^sub>L, 1), j3 + 8 := (\[]\, 1)]" lemma tm25 [transforms_intros]: assumes "ttt = ttt20 + 31 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.." shows "transforms tm25 tps0 ttt tps25" unfolding tm25_def proof (tform tps: tps0 tps24_def tps25_def jk assms) have *: "tps24 ! (j3 + 12) = nltape \" using tps24_def jk by simp then show "clean_tape (tps24 ! (j3 + 12))" using clean_tape_nlcontents by simp have "tps25 = tps24[j3 + 12 := nltape \ |#=| 1]" unfolding tps25_def tps24_def by (simp add: list_update_swap) then show "tps25 = tps24[j3 + 12 := tps24 ! (j3 + 12) |#=| 1]" using * by simp qed definition "tpsE1 \ tps0 [j3 + 11 := (\[]\, 1), j3 + 13 := (\[]\, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := (\\\\<^sub>N\<^sub>L, 1), j3 + 8 := (\[]\, 1)]" lemma tmE1: assumes "ttt = 7 + 2 * nlength (if previous hp1 t = t then 1 else 0)" shows "transforms tmE1 tps25 ttt tpsE1" unfolding tmE1_def proof (tform tps: tps0 tps25_def tpsE1_def jk assms) show "proper_symbols (canrepr (if previous hp1 t = t then 1 else 0))" using proper_symbols_canrepr by simp qed lemma tmE1' [transforms_intros]: assumes "ttt = 9" shows "transforms tmE1 tps25 ttt tpsE1" using tmE1 assms transforms_monotone by (simp add: nlength_le_n) definition "tpsE2 \ tps0 [j3 + 11 := nlltape' (formula_n (relabel \ \')) 0, j3 + 13 := (\[]\, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := (\\\\<^sub>N\<^sub>L, 1), j3 + 8 := (\[]\, 1)]" lemma tmE2 [transforms_intros]: assumes "ttt = 16 + 108 * (nlllength (formula_n \'))\<^sup>2 * (3 + (nllength \)\<^sup>2)" and "previous hp1 t = t" shows "transforms tmE2 tps25 ttt tpsE2" unfolding tmE2_def proof (tform tps: tps0 tpsE1_def tpsE2_def jk assms time: assms(1)) let ?sigma = "[N + (t - Suc 0) * Z..' \ {..[]\\<^sub>N\<^sub>L\<^sub>L, 1)" using tpsE1_def jk nllcontents_Nil by simp show "tpsE1 ! (j3 + 11 + 1) = (\?sigma\\<^sub>N\<^sub>L, 1)" using assms(2) tpsE1_def jk by (simp add: add.commute[of j3 12]) show "tpsE1 ! (j3 + 11 + 2) = (\[]\\<^sub>N\<^sub>L, 1)" using tpsE1_def jk nlcontents_Nil by (simp add: add.commute[of j3 13]) show "tpsE1 ! (j3 + 11 + 3) = (\[]\\<^sub>N\<^sub>L, 1)" using tpsE1_def jk tps0_sym nlcontents_Nil by simp show "tpsE1 ! (j3 + 11 + 4) = (\0\\<^sub>N, 1)" using tpsE1_def jk tps0_sym canrepr_0 by simp show "tpsE1 ! (j3 + 11 + 5) = (\0\\<^sub>N, 1)" using tpsE1_def jk tps0_sym canrepr_0 by simp show "tpsE2 = tpsE1[j3 + 11 := nlltape' (formula_n (relabel ?sigma \')) 0]" unfolding tpsE2_def tpsE1_def using assms(2) by (simp add: list_update_swap[of "j3+11"]) qed definition "tpsTT1 \ tps0 [j3 + 11 := nlltape' (formula_n (relabel \ \)) 0, j3 + 13 := (\[]\, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := (\\\\<^sub>N\<^sub>L, 1), j3 + 8 := (\[]\, 1)]" lemma tmTT1 [transforms_intros]: assumes "ttt = 7 + 108 * (nlllength (formula_n \))\<^sup>2 * (3 + (nllength \)\<^sup>2)" and "previous hp1 t \ t" shows "transforms tmTT1 tps25 ttt tpsTT1" unfolding tmTT1_def proof (tform tps: tps0 tps25_def tpsTT1_def jk time: assms(1)) let ?sigma = "[N + (t - Suc 0) * Z.. t then [N + previous hp1 t * Z.. \ {..[]\\<^sub>N\<^sub>L\<^sub>L, 1)" using tps25_def jk nllcontents_Nil by simp show "tps25 ! (j3 + 11 + 1) = (\?sigma\\<^sub>N\<^sub>L, 1)" using tps25_def jk by (simp add: add.commute[of j3 12]) show "tps25 ! (j3 + 11 + 2) = (\[]\\<^sub>N\<^sub>L, 1)" using tps25_def jk canrepr_0 nlcontents_Nil assms(2) by (simp add: add.commute[of j3 13]) show "tps25 ! (j3 + 11 + 3) = (\[]\\<^sub>N\<^sub>L, 1)" using tps25_def jk tps0_sym nlcontents_Nil by simp show "tps25 ! (j3 + 11 + 4) = (\0\\<^sub>N, 1)" using tps25_def jk tps0_sym canrepr_0 by simp show "tps25 ! (j3 + 11 + 5) = (\0\\<^sub>N, 1)" using tps25_def jk tps0_sym canrepr_0 by simp show "tpsTT1 = tps25[j3 + 11 := nlltape' (formula_n (relabel ?sigma \)) 0]" using tpsTT1_def tps25_def assms(2) canrepr_0 by (simp add: list_update_swap[of "j3+11"]) qed definition "tps26 \ tps0 [j3 + 11 := nlltape' (formula_n (relabel \ (if previous hp1 t = t then \' else \))) 0, j3 + 13 := (\[]\, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := (\\\\<^sub>N\<^sub>L, 1), j3 + 8 := (\[]\, 1)]" lemma nlllength_psi: "nlllength (formula_n \) \ 24 * Z ^ 2 * 2 ^ (4*Z)" proof - let ?V = "3 * Z + G" have "\v. v \ variables \ \ v \ ?V" using psi(1) by auto then have "nlllength (formula_n \) \ fsize \ * Suc (Suc (nlength ?V)) + length \" using nlllength_formula_n by simp also have "... \ fsize \ * Suc (Suc (nlength ?V)) + 2 ^ ?V" using psi by simp also have "... \ ?V * 2 ^ ?V * Suc (Suc (nlength ?V)) + 2 ^ ?V" using psi(2) by (metis add_le_mono1 mult.commute mult_le_mono2) also have "... \ 4*Z * 2 ^ ?V * Suc (Suc (nlength ?V)) + 2 ^ ?V" using Z by simp also have "... \ 4*Z * 2 ^ (4*Z) * Suc (Suc (nlength ?V)) + 2 ^ ?V" proof - have "?V \ 4 * Z" using Z by simp then have "(2::nat) ^ ?V \ 2 ^ (4*Z)" by simp then show ?thesis using add_le_mono le_refl mult_le_mono by presburger qed also have "... \ 4*Z * (2::nat) ^ (4*Z) * Suc (Suc (nlength (4*Z))) + 2 ^ ?V" using Z nlength_mono by simp also have "... \ 4*Z * (2::nat) ^ (4*Z) * Suc (Suc (4*Z)) + 2 ^ ?V" using nlength_le_n by simp also have "... \ 4*Z * 2 ^ (4*Z) * Suc (Suc (4*Z)) + 2 ^ (4*Z)" using Z by simp also have "... \ 4*Z * 2 ^ (4*Z) * (5*Z) + 2 ^ (4*Z)" using Z G by simp also have "... \ 4*Z * 2 ^ (4*Z) * (6*Z)" using Z G by simp also have "... = 24 * Z ^ 2 * 2 ^ (4*Z)" by algebra finally show ?thesis . qed lemma nlllength_psi': "nlllength (formula_n \') \ 24 * Z ^ 2 * 2 ^ (4*Z)" proof - let ?V = "2 * Z + G" have "\v. v \ variables \' \ v \ ?V" using psi'(1) by auto then have "nlllength (formula_n \') \ fsize \' * Suc (Suc (nlength ?V)) + length \'" using nlllength_formula_n by simp also have "... \ fsize \' * Suc (Suc (nlength ?V)) + 2 ^ ?V" using psi' by simp also have "... \ ?V * 2 ^ ?V * Suc (Suc (nlength ?V)) + 2 ^ ?V" using psi'(2) by (metis add_le_mono1 mult.commute mult_le_mono2) also have "... \ 4*Z * 2 ^ ?V * Suc (Suc (nlength ?V)) + 2 ^ ?V" using Z by simp also have "... \ 4*Z * 2 ^ (4*Z) * Suc (Suc (nlength ?V)) + 2 ^ ?V" proof - have "?V \ 4 * Z" using Z by simp then have "(2::nat) ^ ?V \ 2 ^ (4*Z)" by simp then show ?thesis using add_le_mono le_refl mult_le_mono by presburger qed also have "... \ 4*Z * (2::nat) ^ (4*Z) * Suc (Suc (nlength (4*Z))) + 2 ^ ?V" using Z nlength_mono by simp also have "... \ 4*Z * (2::nat) ^ (4*Z) * Suc (Suc (4*Z)) + 2 ^ ?V" using nlength_le_n by simp also have "... \ 4*Z * 2 ^ (4*Z) * Suc (Suc (4*Z)) + 2 ^ (4*Z)" using Z by simp also have "... \ 4*Z * 2 ^ (4*Z) * (5*Z) + 2 ^ (4*Z)" using Z G by simp also have "... \ 4*Z * 2 ^ (4*Z) * (6*Z)" using Z G by simp also have "... = 24 * Z ^ 2 * 2 ^ (4*Z)" by algebra finally show ?thesis . qed lemma tm2526: assumes "ttt = 17 + 108 * (24 * Z ^ 2 * 2 ^ (4*Z))\<^sup>2 * (3 + (nllength \)\<^sup>2)" shows "transforms tm2526 tps25 ttt tps26" unfolding tm2526_def proof (tform) have *: "tps25 ! (j3 + 13) = (\previous hp1 t = t\\<^sub>B, 1)" using tps25_def jk by simp then have **: "(previous hp1 t \ t) = (read tps25 ! (j3 + 13) = \)" using jk tps25_def read_ncontents_eq_0 by force show "read tps25 ! (j3 + 13) = \ \ previous hp1 t \ t" using ** by simp show "read tps25 ! (j3 + 13) \ \ \ previous hp1 t = t" using ** by simp show "read tps25 ! (j3 + 13) = \ \ tps26 = tpsTT1" using tps26_def tpsTT1_def ** by presburger show "read tps25 ! (j3 + 13) \ \ \ tps26 = tpsE2" using tps26_def tpsE2_def ** by presburger let ?tT = "7 + 108 * (nlllength (formula_n \))\<^sup>2 * (3 + (nllength \)\<^sup>2)" let ?tF = "16 + 108 * (nlllength (formula_n \'))\<^sup>2 * (3 + (nllength \)\<^sup>2)" have "?tT + 2 \ 9 + 108 * (24 * Z ^ 2 * 2 ^ (4*Z))\<^sup>2 * (3 + (nllength \)\<^sup>2)" using nlllength_psi linear_le_pow by simp also have "... \ ttt" using assms by simp finally show "?tT + 2 \ ttt" . show "?tF + 1 \ ttt" using nlllength_psi' assms linear_le_pow by simp qed lemma nllength_\: "nllength \ \ 12 * T' * Z^2 + 4 * Z * N" proof - have "n \ N + T' * Z + Z" if "n \ set \" for n proof - consider "n \ set [N + (t - 1) * Z.. set [N + previous hp1 t * Z.. set [N + t * Z.. set [hp0 ! t * G.. set \` by auto then show ?thesis proof (cases) case 1 then have "n \ N + (t - 1) * Z + Z" by simp also have "... \ N + T' * Z + Z" using t(2) by auto finally show ?thesis by simp next case 2 then have "n \ N + (previous hp1 t) * Z + Z" by simp also have "... \ N + t * Z + Z" by (simp add: previous_hp1_le) also have "... \ N + T' * Z + Z" using t(2) by simp finally show ?thesis by simp next case 3 then have "n \ N + t * Z + Z" by simp also have "... \ N + T' * Z + Z" using t(2) by auto finally show ?thesis by simp next case 4 then have "n \ N + (hp0 ! t) * G + G" by simp also have "... \ N + T' * G + G" using t len_hp0 hp0 by simp also have "... \ N + T' * Z + Z" using Z by simp finally show ?thesis by simp qed qed then have "nllength \ \ Suc (N + T' * Z + Z) * (length \)" using nllength_le[of \ "N + T' * Z + Z"] by simp also have "... \ Suc (N + T' * Z + Z) * (3 * Z + G)" proof - have "length \ \ 3 * Z + G" by simp then show ?thesis using mult_le_mono2 by presburger qed also have "... \ Suc (N + T' * Z + Z) * (3 * Z + Z)" using Z by simp also have "... = 4 * Z * Suc (N + T' * Z + Z)" by simp also have "... = 4 * Z + 4 * Z * (N + T' * Z + Z)" by simp also have "... = 4 * Z + 4 * Z * N + 4 * T' * Z^2 + 4 * Z^2" by algebra also have "... \ 4 * Z^2 + 4 * Z * N + 4 * T' * Z^2 + 4 * Z^2" using linear_le_pow by simp also have "... = 8 * Z^2 + 4 * Z * N + 4 * T' * Z^2" by simp also have "... \ 8 * T' * Z^2 + 4 * Z * N + 4 * T' * Z^2" using t by simp also have "... = 12 * T' * Z^2 + 4 * Z * N" by simp finally show ?thesis . qed lemma tm2526' [transforms_intros]: assumes "ttt = 17 + 108 * (24 * Z ^ 2 * 2 ^ (4*Z))\<^sup>2 * (3 + (12 * T' * Z^2 + 4 * Z * N)\<^sup>2)" shows "transforms tm2526 tps25 ttt tps26" proof - have "17 + 108 * (24 * Z ^ 2 * 2 ^ (4*Z))\<^sup>2 * (3 + (nllength \)\<^sup>2) \ ttt" using assms nllength_\ by simp then show ?thesis using tm2526 transforms_monotone by simp qed lemma tm26 [transforms_intros]: assumes "ttt = ttt20 + 31 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. + 17 + 108 * (24 * Z ^ 2 * 2 ^ (4*Z))\<^sup>2 * (3 + (12 * T' * Z^2 + 4 * Z * N)\<^sup>2)" shows "transforms tm26 tps0 ttt tps26" unfolding tm26_def by (tform tps: assms) definition "tps27 \ tps0 [j3 + 11 := nlltape' (formula_n (relabel \ (if previous hp1 t = t then \' else \))) 0, j3 + 13 := (\[]\, 1), j3 + 6 := (\t\\<^sub>N, 1), j3 + 7 := (\0\\<^sub>N, 1), j3 + 12 := (\[]\, 1), j3 + 8 := (\[]\, 1)]" lemma tm27: assumes "ttt = ttt20 + 38 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. + 17 + 108 * (24 * Z ^ 2 * 2 ^ (4*Z))\<^sup>2 * (3 + (12 * T' * Z^2 + 4 * Z * N)\<^sup>2)" shows "transforms tm27 tps0 ttt tps27" unfolding tm27_def proof (tform tps: tps0 tps26_def tps27_def jk) let ?zs = "numlist \" show "tps26 ::: (j3 + 12) = \?zs\" using tps26_def jk nlcontents_def by simp show "proper_symbols ?zs" using proper_symbols_numlist by simp show "ttt = ttt20 + 31 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. + 17 + 108 * (24 * Z\<^sup>2 * 2 ^ (4 * Z))\<^sup>2 * (3 + (12 * T' * Z\<^sup>2 + 4 * Z * N)\<^sup>2) + (tps26 :#: (j3 + 12) + 2 * length (numlist \) + 6)" using tps26_def jk nllength_def assms by simp qed definition "tps27' \ tps0 [j3 + 11 := nlltape' (formula_n (relabel \ (if previous hp1 t = t then \' else \))) 0]" lemma tps27': "tps27 = tps27'" proof - have 1: "tps0[j3 + 13 := (\[]\, Suc 0)] = tps0" using list_update_id[of tps0 "j3+13"] jk tps0 by simp have 2: "tps0[j3 + 12 := (\[]\, Suc 0)] = tps0" using list_update_id[of tps0 "j3+12"] jk tps0 by simp have 3: "tps0[j3 + 8 := (\[]\, Suc 0)] = tps0" using list_update_id[of tps0 "j3+8"] jk tps0 by simp have 4: "tps0[j3 + 7 := (\0\\<^sub>N, Suc 0)] = tps0" using list_update_id[of tps0 "j3+7"] canrepr_0 jk tps0 by simp have 5: "tps0[j3 + 6 := (\t\\<^sub>N, Suc 0)] = tps0" using list_update_id[of tps0 "j3+6"] jk tps0 by simp show ?thesis unfolding tps27_def tps27'_def using tps0 by (simp split del: if_split add: list_update_swap[of _ "j3 + 13"] 1 list_update_swap[of _ "j3 + 12"] 2 list_update_swap[of _ "j3 + 8"] 3 list_update_swap[of _ "j3 + 7"] 4 list_update_swap[of _ "j3 + 6"] 5) qed definition "ttt27 = ttt20 + 38 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. + 17 + 108 * (24 * Z ^ 2 * 2 ^ (4*Z))\<^sup>2 * (3 + (12 * T' * Z^2 + 4 * Z * N)\<^sup>2)" lemma tm27' [transforms_intros]: "transforms tm27 tps0 ttt27 tps27'" using ttt27_def tm27 nllength_\ tps27' transforms_monotone by simp definition "tps28 \ tps0 [1 := nlltape (nss @ formula_n (relabel \ (if previous hp1 t = t then \' else \))), j3 + 11 := (\[]\, 1)]" lemma tm28: assumes "ttt = ttt27 + (11 + 4 * nlllength (formula_n (relabel \ (if previous hp1 t = t then \' else \))))" shows "transforms tm28 tps0 ttt tps28" unfolding tm28_def by (tform tps: tps0(1) tps0 tps27'_def tps28_def jk time: ttt27_def assms) lemma nlllength_relabel_chi_t: "nlllength (formula_n (relabel \ (if previous hp1 t = t then \' else \))) \ Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)" (is "nlllength (formula_n (relabel \ ?phi)) \ _") proof - have "variables ?phi \ {..}" proof (cases "previous hp1 t = t") case True then show ?thesis using psi'(1) by auto next case False then show ?thesis using psi(1) by auto qed then have "nlllength (formula_n (relabel \ ?phi)) \ Suc (nllength \) * nlllength (formula_n ?phi)" using nlllength_relabel_variables by simp moreover have "nlllength (formula_n ?phi) \ 24 * Z ^ 2 * 2 ^ (4*Z)" using nlllength_psi nlllength_psi' by (cases "previous hp1 t = t") simp_all ultimately have "nlllength (formula_n (relabel \ ?phi)) \ Suc (nllength \) * (24 * Z ^ 2 * 2 ^ (4*Z))" by (meson le_trans mult_le_mono2) then show ?thesis by linarith qed definition "tps28' \ tps0 [1 := nlltape (nss @ formula_n (relabel \ (if previous hp1 t = t then \' else \)))]" lemma tps28': "tps28' = tps28" unfolding tps28'_def tps28_def using tps0 list_update_id[of tps0 "j3+11"] by (simp add: list_update_swap[of _ "j3 + 11"]) lemma tm28' [transforms_intros]: assumes "ttt = ttt27 + (11 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)))" shows "transforms tm28 tps0 ttt tps28'" using assms tm28 tps28' nlllength_relabel_chi_t transforms_monotone by simp definition "tps29 \ tps0 [1 := nlltape (nss @ formula_n (relabel \ (if previous hp1 t = t then \' else \))), j3 + 6 := (\Suc t\\<^sub>N, 1)]" lemma tm29 [transforms_intros]: assumes "ttt = ttt27 + 16 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 2 * nlength t" shows "transforms tm29 tps0 ttt tps29" unfolding tm29_def by (tform tps: assms tps0 tps28'_def tps29_def jk) definition "tps30 \ tps0 [1 := nlltape (nss @ formula_n (relabel \ (if previous hp1 t = t then \' else \))), j3 + 6 := (\Suc t\\<^sub>N, 1), j3 + 3 := (\T - 1\\<^sub>N, 1)]" lemma tm30: assumes "ttt = ttt27 + 24 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 2 * nlength t + 2 * nlength T" shows "transforms tm30 tps0 ttt tps30" unfolding tm30_def by (tform tps: assms tps0 tps29_def tps30_def jk) text \ Some helpers for bounding the running time: \ lemma pow4_le_2pow4: "z ^ 4 \ 2 ^ (4*z)" for z :: nat proof - have "z ^ 4 = (z ^ 2) ^ 2" by simp also have "... \ (2^(2*z)) ^ 2" using pow2_le_2pow2 power_mono by blast also have "... = 2^(4*z)" by (metis mult.commute mult_2_right numeral_Bit0 power_mult) finally show ?thesis . qed lemma pow8_le_2pow8: "z ^ 8 \ 2 ^ (8*z)" for z :: nat proof - have "z ^ 8 = (z ^ 2) ^ 4" by simp also have "... \ (2^(2*z)) ^ 4" using pow2_le_2pow2 power_mono by blast also have "... = 2^(8*z)" by (metis mult.commute mult_2_right numeral_Bit0 power_mult) finally show ?thesis . qed lemma Z_sq_le: "Z^2 \ 2^(16*Z)" proof - have "Z^2 \ 2^(2*Z)" using pow2_le_2pow2[of Z] by simp also have "... \ 2^(16*Z)" by simp finally show "Z^2 \ 2^(16*Z)" . qed lemma time_bound: "ttt27 + 24 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 2 * nlength t + 2 * nlength T \ 16114765 * 2^(16*Z) * N^6" proof - have sum_sq: "(a + b) ^ 2 \ Suc (2 * b) * a ^ 2 + b ^ 2" for a b :: nat proof - have "(a + b) ^ 2 = a ^ 2 + 2 * a * b + b ^ 2" by algebra also have "... \ a ^ 2 + 2 * a ^ 2 * b + b ^ 2" using linear_le_pow by simp also have "... = Suc (2 * b) * a ^ 2 + b ^ 2" by simp finally show ?thesis . qed have 1: "t < N" using t T' by simp then have 15: "t \ N" by simp have 2: "nlength t < N" using 1 nlength_le_n dual_order.strict_trans2 by blast have 25: "nlength T \ N" using T T' nlength_le_n by (meson le_trans order_less_imp_le) have 27: "nlength (t - 1) < N" using t(1) nlength_mono 2 by (metis diff_less dual_order.strict_trans2 less_numeral_extra(1) linorder_not_less order.asym) have 3: "t * Z < N * Z" using 1 Z_ge_1 by simp then have 4: "N + t * Z < Suc Z * N" using 1 by simp have 41: "N + t * Z + Z \ Suc Z * N" proof - have "N + t * Z + Z \ N + (N - 1) * Z + Z" using 1 N by auto then show ?thesis using N by (metis One_nat_def Suc_n_not_le_n ab_semigroup_add_class.add_ac(1) add.commute mult.commute mult_eq_if times_nat.simps(2)) qed have 42: "(t + Z)^2 \ (N + Z) ^ 2" using 1 by simp have 45: "(N + Z) ^ 2 \ 3*Z*N^2 + Z^2" proof - have "(N + Z) ^ 2 \ Suc (2 * Z) * N^2 + Z^2" using sum_sq by simp also have "... \ 3*Z * N^2 +Z^2" using Z_ge_1 by simp finally show ?thesis . qed have 5: "nlength (previous hp1 t) < N" using previous_hp1_le 1 by (meson dual_order.strict_trans2 nlength_le_n) then have 51: "nlength (previous hp1 t) \ N" by simp have 6: "nlength (N + t * Z) < Suc Z * N" using 4 nlength_le_n by (meson le_trans linorder_not_le) have "nllength \ \ 12 * N * Z^2 + 4 * Z * N" proof - have "nllength \ \ 12 * T' * Z^2 + 4 * Z * N" using nllength_\ by simp also have "... \ 12 * N * Z^2 + 4 * Z * N" using T' by simp finally show ?thesis . qed have 7: "previous hp1 t \ N" using previous_hp1_le 15 by simp have 65: "(nlength (previous hp1 t) + nlength Z)\<^sup>2 < (N + Z) ^ 2" proof - have "nlength (previous hp1 t) + nlength Z < N + Z" using 7 2 5 by (simp add: add_less_le_mono nlength_le_n) then show ?thesis by (simp add: power_strict_mono) qed have 66: "N + previous hp1 t * Z + Z \ Suc Z * N" using 41 previous_hp1_le by (meson add_le_mono le_trans less_or_eq_imp_le mult_le_mono) have 67: "(nlength t + nlength Z)\<^sup>2 \ 3 * Z * N^2 + Z^2" proof - have "nlength t + nlength Z < N + Z" using nlength_le_n 2 by (simp add: add_less_le_mono) then have "(nlength t + nlength Z)^2 < (N + Z) ^ 2" by (simp add: power_strict_mono) then show ?thesis using 45 by simp qed have "nlength (previous hp1 t * Z) \ N * Z" using nlength_le_n previous_hp1_le 1 by (meson le_trans less_or_eq_imp_le mult_le_mono) have 75: "max (nlength N) (nlength (t * Z)) \ Suc Z * N" proof - have "max (nlength N) (nlength (t * Z)) \ nlength (max N (t * Z))" using max_nlength by simp also have "... \ nlength (N + t * Z)" by (simp add: nlength_mono) finally show ?thesis using 6 by simp qed then have 78: "max (nlength N) (nlength (previous hp1 t * Z)) \ Suc Z * N" using previous_hp1_le nlength_mono by (smt (verit, best) Groups.mult_ac(2) le_trans max_def mult_le_mono2) have 79: "nlength (N + t * Z + Z) \ Suc Z * N + Z" proof - have "N + t * Z + Z \ Suc Z * N + Z" using previous_le 15 by simp then show ?thesis using nlength_le_n le_trans by blast qed have 8: "nllength [N + previous hp1 t * Z.. 2 * Z^2 * N + 2 * Z^2" proof - have "nllength [N + previous hp1 t * Z.. Suc (nlength (N + previous hp1 t * Z + Z)) * Z" using nllength_upt_le_len_mult_max by (metis diff_add_inverse) moreover have "nlength (N + previous hp1 t * Z + Z) \ Suc Z * N + Z" proof - have "N + previous hp1 t * Z + Z \ Suc Z * N + Z" using previous_le 15 7 by simp then show ?thesis using nlength_le_n le_trans by blast qed ultimately have "nllength [N + previous hp1 t * Z.. Suc (Suc Z * N + Z) * Z" by (meson Suc_le_mono le_trans less_or_eq_imp_le mult_le_mono) also have "... = (Z^2 + Z) * Suc N" by (metis add.commute mult.commute mult.left_commute mult_Suc nat_arith.suc1 power2_eq_square) also have "... \ (Z^2 + Z^2) * Suc N" by (meson add_le_cancel_left linear_le_pow mult_le_mono1 rel_simps(51)) also have "... = 2 * Z^2 * Suc N" by simp also have "... = 2 * Z^2 * N + 2 * Z^2" by simp finally show ?thesis . qed have 84: "Z * Suc Z \ 2 * Z^2" by (simp add: power2_eq_square) have 85: "nlength (N + previous hp1 t * Z) \ Suc Z * N" proof - have "nlength (N + previous hp1 t * Z) \ nlength (N + t * Z)" using previous_hp1_le nlength_mono by simp then show ?thesis using 6 by simp qed have 9: "Suc Z \ 2*Z" using Z_ge_1 by simp then have 91: "Suc Z ^ 2 \ 4 * Z^2" by (metis mult_2_right numeral_Bit0 power2_eq_square power2_nat_le_eq_le power_mult_distrib) have 99: "Z^2 \ 81" proof - have "Z * Z \ 9 * 9" using Z_ge_9 mult_le_mono by presburger moreover have "9 * 9 = (81::nat)" by simp ultimately show ?thesis by (simp add: power2_eq_square) qed have part1: "ttt8 \ 241 * Z^2 + 266 * Z^2 * N ^ 6" proof - have "ttt8 \ 168 + 153 * N ^ 6 + 5 * t + 26 * (t + Z)\<^sup>2 + 47 * Z + 15 * Z * (N + t * Z)" using ttt8_def T' by simp also have "... \ 168 + 153 * N ^ 6 + 5 * N + 26 * (t + Z)\<^sup>2 + 47 * Z + 15 * Z * (N + t * Z)" using 15 by simp also have "... \ 168 + 153 * N ^ 6 + 5 * N + 26 * (N + Z)\<^sup>2 + 47 * Z + 15 * Z * (N + t * Z)" using 42 by simp also have "... \ 168 + 153 * N ^ 6 + 5 * N + 26 * (3*Z*N^2 + Z^2) + 47 * Z + 15 * Z * (N + t * Z)" using 45 by simp also have "... \ 168 + 153 * N ^ 6 + 5 * N + 26 * (3*Z*N^2 + Z^2) + 47 * Z + 15 * Z * Suc Z * N" using 4 by (metis (mono_tags, lifting) add_left_mono less_or_eq_imp_le mult.assoc mult_le_mono2) also have "... \ 168 + 153 * N ^ 6 + 5 * N + 26 * (3*Z*N^2 + Z^2) + 47 * Z + 30 * Z^2 * N" using `Z * Suc Z \ 2 * Z^2` by simp also have "... = 168 + 153 * N ^ 6 + 5 * N + 78 * Z * N^2 + 26*Z^2 + 47 * Z + 30 * Z^2 * N" by simp also have "... \ 168 + 158 * N ^ 6 + 78 * Z * N^2 + 73 * Z^2 + 30 * Z^2 * N" using linear_le_pow[of 6 N] linear_le_pow[of 2 Z] by simp also have "... \ 168 + 158 * N ^ 6 + 78 * Z^2 * N^2 + 73 * Z^2 + 30 * Z^2 * N^2" using linear_le_pow by (metis add_le_mono add_le_mono1 le_square mult_le_mono1 mult_le_mono2 nat_add_left_cancel_le power2_eq_square) also have "... = 168 + 158 * N ^ 6 + 108 * Z^2 * N^2 + 73 * Z^2" by simp also have "... \ 168*Z^2 + 158 * N ^ 6 + 108 * Z^2 * N^2 + 73 * Z^2" using Z_ge_1 by simp also have "... \ 241 * Z^2 + 158 * N ^ 6 + 108 * Z^2 * N^6" using pow_mono'[of 2 6 N] by simp also have "... \ 241 * Z^2 + 158 * Z^2 * N ^ 6 + 108 * Z^2 * N^6" using Z_ge_1 by simp also have "... = 241 * Z^2 + 266 * Z^2 * N ^ 6" by simp finally show ?thesis . qed have part2: "ttt10 - ttt8 \ 63 * Z^2 + 226 * Z^2 * N^6" proof - have "ttt10 - ttt8 = 37 + 26 * (nlength (previous hp1 t) + nlength Z)\<^sup>2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 4 * nllength [N + previous hp1 t * Z.. 37 + 26 * (nlength (previous hp1 t) + nlength Z)\<^sup>2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 4 * (2 * Z^2 * N + 2 * Z^2) + 2 * nlength (N + previous hp1 t * Z)" using 8 by simp also have "... \ 37 + 26 * (N + Z)\<^sup>2 + 3 * max (nlength N) (nlength (previous hp1 t * Z)) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 4 * (2 * Z^2 * N + 2 * Z^2) + 2 * nlength (N + previous hp1 t * Z)" using 65 by linarith also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 8 * Z^2 * N + 8 * Z^2 + 2 * nlength (N + previous hp1 t * Z)" using 78 45 by auto also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * nlength (N + previous hp1 t * Z + Z)) + 8 * Z^2 * N + 8 * Z^2 + 2 * Suc Z * N" using 85 by linarith also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * nlength (Suc Z * N)) + 8 * Z^2 * N + 8 * Z^2 + 2 * Suc Z * N" using 66 nlength_mono add_le_mono le_refl mult_le_mono by presburger also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * (Suc Z * N)) + 8 * Z^2 * N + 8 * Z^2 + 2 * Suc Z * N" using nlength_le_n add_le_mono le_refl mult_le_mono by presburger also have "... = 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + 43 * Suc Z + Suc Z * 9 * Suc Z * N + 8 * Z^2 * N + 8 * Z^2 + 2 * Suc Z * N" by algebra also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + 43 * 2 * Z + 2 * Z * 9 * Suc Z * N + 8 * Z^2 * N + 8 * Z^2 + 2 * Suc Z * N" using 9 by (simp add: add_le_mono) also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + 86 * Z + 36 * Z * Z * N + 8 * Z^2 * N + 8 * Z^2 + 2 * Suc Z * N" by simp also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + 86 * Z + 44 * Z^2 * N + 8 * Z^2 + 2 * Suc Z * N" by (simp add: power2_eq_square) also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + 86 * Z + 44 * Z^2 * N + 8 * Z^2 + 4 * Z * N" using 9 by simp also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * (Suc Z * N) + 86 * Z + 48 * Z^2 * N + 8 * Z^2" using linear_le_pow by simp also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 3 * 2 * Z * N + 86 * Z + 48 * Z^2 * N + 8 * Z^2" using 9 by simp also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 6 * Z * N + 86 * Z^2 * N + 48 * Z^2 * N + 8 * Z^2" using linear_le_pow[of 2 Z] N by simp (metis N le_trans mult_le_mono1 nat_mult_1) also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 140 * Z^2 * N + 8 * Z^2" using linear_le_pow[of 2 Z] by simp also have "... \ 37 + 26 * (3*Z*N^2 + Z^2) + 148 * Z^2 * N" using N by simp also have "... = 37 + 78 * Z * N^2 + 26 * Z^2 + 148 * Z^2 * N" by simp also have "... \ 63 * Z^2 + 78 * Z * N^2 + 148 * Z^2 * N" using Z_ge_1 by simp also have "... \ 63 * Z^2 + 78 * Z^2 * N^2 + 148 * Z^2 * N" using linear_le_pow by simp also have "... \ 63 * Z^2 + 226 * Z^2 * N^2" using linear_le_pow by simp also have "... \ 63 * Z^2 + 226 * Z^2 * N^6" using pow_mono'[of 2 6 N] by simp finally show ?thesis . qed have 10: "nllength hp0 \ N ^ 2" proof - have "\n\set hp0. n \ T'" using hp0 by (metis in_set_conv_nth) then have "nllength hp0 \ Suc T' * Suc T'" using nllength_le[of hp0 T'] len_hp0 by simp also have "... \ N * N" using T' Suc_leI mult_le_mono by presburger also have "... = N ^ 2" by algebra finally show ?thesis . qed have 11: "nllength [N + t * Z.. 2 * Z^2 * N + Z" proof - have "nllength [N + t * Z.. Suc (N + t * Z + Z) * Z" using nllength_upt_le[of "N + t * Z" "N + t * Z + Z"] by simp also have "... \ Suc (Suc Z * N) * Z" using 41 by simp also have "... = (Z*Z + Z)*N + Z" by (metis add.commute mult.commute mult.left_commute mult_Suc) also have "... \ 2 * Z^2 * N + Z" using linear_le_pow[of 2 Z] by (simp add: power2_eq_square) finally show ?thesis . qed have 12: "nlength (hp0 ! t) + nlength G \ N + Z" proof - have "nlength (hp0 ! t) + nlength G \ hp0 ! t + G" using nlength_le_n by (simp add: add_mono) also have "... \ T' + Z" using Z by (simp add: add_le_mono hp0 le_imp_less_Suc len_hp0 t(2)) also have "... \ N + Z" using T' by simp finally show ?thesis . qed have part3: "ttt20 - ttt10 \ 120 * Z^2 + 206 * Z^2 * N ^ 4" proof - have "ttt20 - ttt10 = 80 + 2 * nlength (t - 1) + 26 * (nlength t + nlength Z)\<^sup>2 + 3 * max (nlength N) (nlength (t * Z)) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z) + 26 * (nlength (hp0 ! t) + nlength G)\<^sup>2" using ttt20_def ttt10_def by simp also have "... \ 80 + 2 * N + 26 * (3 * Z * N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * nlength (N + t * Z + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z) + 26 * (nlength (hp0 ! t) + nlength G)\<^sup>2" using 27 67 75 by auto also have "... \ 80 + 2 * N + 26 * (3 * Z * N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * (Suc Z * N + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z) + 26 * (nlength (hp0 ! t) + nlength G)\<^sup>2" using 79 add_le_mono le_refl mult_le_mono by presburger also have "... \ 80 + 2 * N + 26 * (3 * Z * N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * (Suc Z * N + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z) + 26 * (N + Z)\<^sup>2" using 51 12 by (metis add_le_cancel_right add_le_mono nat_add_left_cancel_le nat_mult_le_cancel_disj power2_nat_le_eq_le) also have "... \ 80 + 4 * N + 52 * (3 * Z * N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * (Suc Z * N + Z)) + 4 * nllength [N + t * Z..2 + 2 * nlength (N + t * Z)" using 45 by simp also have "... \ 80 + 4 * N + 52 * (3 * Z * N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * (Suc Z * N + Z)) + 4 * nllength [N + t * Z..2 + 2 * Suc Z * N" using 6 by linarith also have "... \ 80 + 4 * N + 52 * (3 * Z * N^2 + Z^2) + 3 * (Suc Z * N) + Suc Z * (43 + 9 * (Suc Z * N + Z)) + 4 * (2 * Z^2 * N + Z) + 21 * (N^2)\<^sup>2 + 2 * Suc Z * N" using 11 10 add_le_mono le_refl mult_le_mono2 power2_nat_le_eq_le by presburger also have "... = 80 + 4 * N + 156 * Z * N^2 + 52 * Z^2 + 3 * (Suc Z * N) + Suc Z * (43 + 9 * (Suc Z * N + Z)) + 8 * Z^2 * N + 4 * Z + 21 * N ^ 4 + 2 * Suc Z * N" by simp also have "... = 80 + 156 * Z * N^2 + 52 * Z^2 + Suc Z * 43 + Suc Z * 9 * Suc Z * N + Suc Z * 9 * Z + (4 + 5 * Suc Z + 8 * Z^2) * N + 4 * Z + 21 * N ^ 4" by algebra also have "... = 123 + 156 * Z * N^2 + 52 * Z^2 + 47 * Z + 9 * Suc Z * Suc Z * N + Suc Z * 9 * Z + (9 + 5 * Z + 8 * Z^2) * N + 21 * N ^ 4" by simp also have "... = 123 + 156 * Z * N^2 + 52 * Z^2 + 47 * Z + 9 * Z * Suc Z + (9 * Suc Z^2 + 9 + 5 * Z + 8 * Z^2) * N + 21 * N ^ 4" by algebra also have "... \ 123 + 156 * Z * N^2 + 52 * Z^2 + 47 * Z + 9 * Z * Suc Z + (44 * Z^2 + 9 + 5 * Z) * N + 21 * N ^ 4" using 91 by simp also have "... \ 123 + 156 * Z * N^2 + 70 * Z^2 + 47 * Z + (44 * Z^2 + 9 + 5 * Z) * N + 21 * N ^ 4" using 84 by simp also have "... \ 123 + 156 * Z * N^2 + 70 * Z^2 + 47 * Z^2 + (44 * Z^2 + 9 + 5 * Z^2) * N + 21 * N ^ 4" using linear_le_pow[of 2 Z] add_le_mono le_refl mult_le_mono power2_nat_le_imp_le by presburger also have "... \ 123 + 156 * Z * N^2 + 117 * Z^2 + (49 * Z^2 + 9) * N ^ 4 + 21 * N ^ 4" using linear_le_pow[of 4 N] by simp also have "... \ 123 + 156 * Z * N^4 + 117 * Z^2 + (49 * Z^2 + 9) * N ^ 4 + 21 * N ^ 4" using pow_mono'[of 2 4 N] by simp also have "... = 123 + 117 * Z^2 + (156 * Z + 49 * Z^2 + 30) * N ^ 4" by algebra also have "... \ 123 + 117 * Z^2 + (205 * Z^2 + 30) * N ^ 4" using linear_le_pow[of 2 Z] by simp also have "... \ 120 * Z^2 + (205 * Z^2 + 30) * N ^ 4" using 99 by simp also have "... \ 120 * Z^2 + 206 * Z^2 * N ^ 4" using 99 by simp finally show ?thesis . qed have 12: "hp0 ! t * G + G \ Z * N" proof - have "hp0 ! t * G + G \ T' * G + G" using hp0 t(2) len_hp0 by simp also have "... \ (N - 1) * G + G" using T' by auto also have "... = N * G" using T' by (metis Suc_diff_1 add.commute less_nat_zero_code mult_Suc not_gr_zero) also have "... \ Z * N" using Z by simp finally show ?thesis . qed have 13: "Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) \ 43 * Z + 9 * Z^2 * N" proof - have "Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) \ Suc G * (43 + 9 * (hp0 ! t * G + G))" using nlength_le_n add_le_mono le_refl mult_le_mono by presburger also have "... \ Suc G * (43 + 9 * (Z * N))" using 12 add_le_mono less_or_eq_imp_le mult_le_mono by presburger also have "... \ Z * (43 + 9 * (Z * N))" using G Z by simp also have "... = 43 * Z + 9 * N * Z * Z" by algebra also have "... = 43 * Z + 9 * Z^2 * N" by algebra finally show ?thesis . qed have 14: "nllength [hp0 ! t * G.. Z^2 * N" proof - have "nllength [hp0 ! t * G.. (hp0 ! t * G + G) * Suc G" using nllength_upt_le[of "hp0 ! t * G" "hp0 ! t * G + G"] by simp also have "... \ (hp0 ! t * G + G) * Z" using Z G by simp also have "... \ N * Z * Z" using 12 by (simp add: mult.commute) also have "... = Z^2 * N" by algebra finally show ?thesis . qed have 15: "nlength (hp0 ! t) \ N" using T' hp0 t(2) len_hp0 nlength_le_n by (metis le_imp_less_Suc le_trans less_or_eq_imp_le) have 16: "nlength (hp0 ! t * G) \ Z * N" proof - have "nlength (hp0 ! t * G) \ hp0 ! t * G" using nlength_le_n by simp also have "... \ T' * G" using Z T' hp0 t(2) len_hp0 by simp also have "... \ Z * N" using Z T' by simp finally show ?thesis . qed have 17: "(12 * T' * Z\<^sup>2 + 4 * Z * N) ^ 2 \ 256 * Z^4 * N^2" proof - have "(12 * T' * Z\<^sup>2 + 4 * Z * N) ^ 2 \ (12 * N * Z^2 + 4 * Z * N)^2" using T' by simp also have "... \ (12 * N * Z^2 + 4 * Z^2 * N)^2" using linear_le_pow[of 2 Z] add_le_mono le_refl mult_le_mono power2_nat_le_eq_le power2_nat_le_imp_le by presburger also have "... = 256 * Z^4 * N^2" by algebra finally show ?thesis . qed have 18: "108 * (24 * Z\<^sup>2 * 2 ^ (4 * Z))\<^sup>2 * (3 + (12 * T' * Z\<^sup>2 + 4 * Z * N)\<^sup>2) \ 16111872 * 2^(16*Z) * N^2" proof - have "108 * (24 * Z\<^sup>2 * 2 ^ (4 * Z))\<^sup>2 = 62208 * (Z^2 * 2 ^ (4*Z))^2" by algebra also have "... = 62208 * Z^(2*2) * 2 ^ (2*(4*Z))" by (metis (no_types, lifting) mult.assoc power_even_eq power_mult_distrib) also have "... = 62208 * Z^4 * 2 ^ (8*Z)" by simp finally have *: "108 * (24 * Z\<^sup>2 * 2 ^ (4 * Z))\<^sup>2 = 62208 * Z^4 * 2 ^ (8*Z)" . have "3 + (12 * T' * Z\<^sup>2 + 4 * Z * N)\<^sup>2 \ 3 + 256 * Z^4 * N^2" using 17 by simp moreover have "Z^4 * N^2 \ 1" using Z_ge_1 N by simp ultimately have "3 + (12 * T' * Z\<^sup>2 + 4 * Z * N)\<^sup>2 \ 259 * Z^4 * N^2" by linarith then have "108 * (24 * Z\<^sup>2 * 2 ^ (4 * Z))\<^sup>2 * (3 + (12 * T' * Z\<^sup>2 + 4 * Z * N)\<^sup>2) \ 16111872 * Z^4 * 2 ^ (8*Z) * Z^4 * N^2" using * by simp also have "... = 16111872 * Z^8 * 2 ^ (8*Z) * N^2" by simp also have "... \ 16111872 * 2^(8*Z) * 2 ^ (8*Z) * N^2" using pow8_le_2pow8 by simp also have "... = 16111872 * 2^(8*Z+8*Z) * N^2" by (metis (no_types, lifting) mult.assoc power_add) also have "... = 16111872 * 2^(16*Z) * N^2" by simp finally show ?thesis . qed have 19: "nllength \ \ 16 * Z^2 * N" proof - have "nllength \ \ 12 * T' * Z^2 + 4 * Z * N" using nllength_\ by simp also have "... \ 12 * N * Z^2 + 4 * Z * N" using T' by simp also have "... \ 12 * Z^2 * N + 4 * Z^2 * N" using linear_le_pow[of 2 Z] by simp also have "... \ 16 * Z^2 * N" by simp finally show ?thesis . qed have part4: "ttt27 - ttt20 \ 50 * Z + 16111936 * 2^(16*Z) * N^2" proof - have "ttt27 - ttt20 = 55 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G.. + 108 * (24 * Z\<^sup>2 * 2 ^ (4 * Z))\<^sup>2 * (3 + (12 * T' * Z\<^sup>2 + 4 * Z * N)\<^sup>2)" using ttt27_def ttt20_def by linarith also have "... \ 55 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 4 * nllength [hp0 ! t * G..2 * 2 ^ (4 * Z))\<^sup>2 * (3 + (12 * T' * Z\<^sup>2 + 4 * Z * N)\<^sup>2)" using 19 by (simp add: mult.commute) also have "... \ 55 + Suc G * (43 + 9 * nlength (hp0 ! t * G + G)) + 2 * Z * N + 2 * N + 52 * Z^2 * N + 16111872 * 2^(16*Z) * N^2" using 14 15 16 18 by linarith also have "... \ 55 + 43 * Z + 9 * Z^2 * N + 2 * Z * N + 2 * N + 52 * Z^2 * N + 16111872 * 2^(16*Z) * N^2" using 13 by linarith also have "... = 55 + 43 * Z + 2 * Z * N + 2 * N + 61 * Z^2 * N + 16111872 * 2^(16*Z) * N^2" by simp also have "... \ 50 * Z + 2 * Z * N + 2 * N + 61 * Z^2 * N + 16111872 * 2^(16*Z) * N^2" using Z_ge_9 by simp also have "... \ 50 * Z + 3 * Z * N + 61 * Z^2 * N + 16111872 * 2^(16*Z) * N^2" using Z_ge_9 by simp also have "... \ 50 * Z + 64 * Z^2 * N + 16111872 * 2^(16*Z) * N^2" using linear_le_pow[of 2 Z] by simp also have "... \ 50 * Z + 64 * Z^2 * N^2 + 16111872 * 2^(16*Z) * N^2" using linear_le_pow[of 2 N] by simp also have "... \ 50 * Z + 64 * 2^(2*Z) * N^2 + 16111872 * 2^(16*Z) * N^2" using pow2_le_2pow2 by simp also have "... \ 50 * Z + 64 * 2^(16*Z) * N^2 + 16111872 * 2^(16*Z) * N^2" by simp also have "... = 50 * Z + 16111936 * 2^(16*Z) * N^2" by simp finally show ?thesis . qed have part5: "24 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 2 * nlength t + 2 * nlength T \ 24 + 1633 * 2^(8*Z) * N" proof - have "24 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 2 * nlength t + 2 * nlength T \ 24 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 4 * N" using 25 2 by simp also have "... \ 24 + 4 * (Suc (16 * Z^2 * N) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 4 * N" using 19 by (simp add: mult.commute) also have "... \ 24 + 1632 * Z^2 * N * Z ^ 2 * 2 ^ (4*Z) + 4 * N" using Z N by simp also have "... = 24 + 1632 * Z^4 * 2 ^ (4*Z) * N + 4 * N" by simp also have "... \ 24 + 1632 * 2^(4*Z) * 2 ^ (4*Z) * N + 4 * N" using pow4_le_2pow4 by simp also have "... = 24 + 1632 * 2^(8*Z) * N + 4 * N" by (metis (no_types, lifting) ab_semigroup_mult_class.mult_ac(1) add_mult_distrib numeral_Bit0 power_add) also have "... \ 24 + 1632 * 2^(8*Z) * N + 2^(8*Z) * N" proof - have "(4::nat) \ 2 ^ 8" by simp also have "... \ 2 ^ (8*Z)" using Z_ge_1 by (metis nat_mult_1_right nat_mult_le_cancel_disj one_le_numeral power_increasing) finally have "(4::nat) \ 2 ^ (8*Z)" . then show ?thesis by simp qed also have "... \ 24 + 1633 * 2^(8*Z) * N" by simp finally show ?thesis . qed show "ttt27 + 24 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 2 * nlength t + 2 * nlength T \ 16114765 * 2^(16*Z) * N^6" proof - have "ttt27 \ ttt20 + 50 * Z + 16111936 * 2^(16*Z) * N^2" using part4 ttt27_def by simp also have "... \ ttt10 + 120 * Z^2 + 206 * Z^2 * N ^ 4 + 50 * Z + 16111936 * 2^(16*Z) * N^2" using part3 ttt20_def by simp also have "... \ ttt8 + 63 * Z^2 + 226 * Z^2 * N^6 + 120 * Z^2 + 206 * Z^2 * N ^ 4 + 50 * Z + 16111936 * 2^(16*Z) * N^2" using part2 ttt10_def by simp also have "... \ 241 * Z^2 + 266 * Z^2 * N ^ 6 + 63 * Z^2 + 226 * Z^2 * N^6 + 120 * Z^2 + 206 * Z^2 * N ^ 4 + 50 * Z + 16111936 * 2^(16*Z) * N^2" using part1 by simp also have "... = 424 * Z^2 + 492 * Z^2 * N ^ 6 + 206 * Z^2 * N ^ 4 + 50 * Z + 16111936 * 2^(16*Z) * N^2" by simp also have "... \ 474 * Z^2 + 492 * Z^2 * N ^ 6 + 206 * Z^2 * N ^ 4 + 16111936 * 2^(16*Z) * N^2" using linear_le_pow[of 2 Z] by simp also have "... \ 474 * Z^2 + 698 * Z^2 * N ^ 6 + 16111936 * 2^(16*Z) * N^2" using pow_mono'[of 4 6 N] by simp also have "... \ 474 * Z^2 + 698 * Z^2 * N ^ 6 + 16111936 * 2^(16*Z) * N^6" using pow_mono'[of 2 6 N] by simp also have "... \ 474 * Z^2 + 698 * 2^(16*Z) * N ^ 6 + 16111936 * 2^(16*Z) * N^6" using Z_sq_le by simp also have "... = 474 * Z^2 + 16112634 * 2^(16*Z) * N^6" by simp also have "... \ 474 * 2^(16*Z) + 16112634 * 2^(16*Z) * N^6" using Z_sq_le by simp also have "... \ 16113108 * 2^(16*Z) * N^6" using N by simp finally have "ttt27 \ 16113108 * 2^(16*Z) * N^6" . then have "ttt27 + 24 + 4 * (Suc (nllength \) * 24 * Z ^ 2 * 2 ^ (4*Z)) + 2 * nlength t + 2 * nlength T \ 16113108 * 2^(16*Z) * N^6 + 24 + 1633 * 2^(8*Z) * N" using part5 by simp also have "... \ 16113108 * 2^(16*Z) * N^6 + 24 + 1633 * 2^(16*Z) * N" by simp also have "... \ 16113108 * 2^(16*Z) * N^6 + 24 + 1633 * 2^(16*Z) * N^6" using linear_le_pow[of 6 N] by simp also have "... = 24 + 16114741 * 2^(16*Z) * N^6" by simp also have "... \ 24 * 2^(16*Z) + 16114741 * 2^(16*Z) * N^6" using Z_sq_le by simp also have "... \ 16114765 * 2^(16*Z) * N^6" using N by simp finally show ?thesis . qed qed lemma tm30': assumes "ttt = 16114765 * 2^(16*Z) * N^6" shows "transforms tm30 tps0 ttt tps30" using tm30 time_bound transforms_monotone assms by simp end (* context tps0 *) end (* locale turing_machine_chi *) lemma transforms_tm_chi: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and k G N Z T' T t :: nat and hp0 hp1 :: "nat list" and \ \' :: formula fixes nss :: "nat list list" assumes "length tps = k" and "1 < j1" "j1 < j2" "j2 < j3" "j3 + 17 < k" and "G \ 3" and "Z = 3 * G" and "N \ 1" and "length hp0 = Suc T'" and "\i T'" and "length hp1 = Suc T'" and "\i T'" and "0 < t" "t \ T'" and "0 < T" "T \ T'" and "T' < N" and "variables \ \ {..<3*Z+G}" "fsize \ \ (3*Z+G) * 2 ^ (3*Z+G)" "length \ \ 2 ^ (3*Z+G)" and "variables \' \ {..<2*Z+G}" "fsize \' \ (2*Z+G) * 2 ^ (2*Z+G)" "length \' \ 2 ^ (2*Z+G)" assumes "tps ! 1 = nlltape nss" "tps ! j1 = (\hp0\\<^sub>N\<^sub>L, 1)" "tps ! j2 = (\hp1\\<^sub>N\<^sub>L, 1)" "tps ! j3 = (\N\\<^sub>N, 1)" "tps ! (j3 + 1) = (\G\\<^sub>N, 1)" "tps ! (j3 + 2) = (\Z\\<^sub>N, 1)" "tps ! (j3 + 3) = (\T\\<^sub>N, 1)" "tps ! (j3 + 4) = (\formula_n \\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps ! (j3 + 5) = (\formula_n \'\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps ! (j3 + 6) = (\t\\<^sub>N, 1)" "\i. 6 < i \ i < 17 \ tps ! (j3 + i) = (\[]\, 1)" assumes "tps' = tps [1 := nlltape (nss @ formula_n (relabel ([N + (t - 1) * Z.. t then [N + previous hp1 t * Z..' else \))), j3 + 6 := (\Suc t\\<^sub>N, 1), j3 + 3 := (\T - 1\\<^sub>N, 1)]" assumes "ttt = 16114765 * 2 ^ (16*Z) * N^6" shows "transforms (tm_chi j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_chi j1 j2 j3 . show ?thesis using loc.tm30'[OF assms(1-34)] loc.tps30_def[OF assms(1-34)] assms(35,36) loc.tm30_eq_tm_chi by simp qed subsubsection \A Turing machine for $\Phi_9$ proper\ text \ The formula $\Phi_9$ is a conjunction of formulas $\chi_t$. The TM @{const tm_chi} decreases the number on tape $j_3 + 3$. If this tape is initialized with $T'$, then a while loop with @{const tm_chi} as its body will generate $\Phi_9$. The next TM is such a machine: \ definition tm_PHI9 :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_PHI9 j1 j2 j3 \ WHILE [] ; \rs. rs ! (j3 + 3) \ \ DO tm_chi j1 j2 j3 DONE" lemma tm_PHI9_tm: assumes "0 < j1" and "j1 < j2" and "j2 < j3" and "j3 + 17 < k" and "G \ 6" shows "turing_machine k G (tm_PHI9 j1 j2 j3)" unfolding tm_PHI9_def using tm_chi_tm turing_machine_loop_turing_machine Nil_tm assms by simp lemma map_nth: fixes zs ys f n i assumes "zs = map f [0..t. formula_n (\ t)) [0..t. \ t) [0.. b" and "b \ c" shows "[a.. The semantics of the TM @{const tm_PHI9} can be proved inside the locale @{locale reduction_sat_x} because it is a fairly simple TM. \ context reduction_sat_x begin text \ The TM @{const tm_chi} is the first TM whose semantics we transfer into the locale @{locale reduction_sat_x}. \ lemma tm_chi: fixes tps0 :: "tape list" and k t' t :: nat and j1 j2 j3 :: tapeidx fixes nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j1" "j1 < j2" "j2 < j3" "j3 + 17 < k" and t: "0 < t" "t \ T'" and T: "0 < t'" "t' \ T'" assumes "hp0 = map (\t. exc (zeros m) t <#> 0) [0..t. exc (zeros m) t <#> 1) [0..hp0\\<^sub>N\<^sub>L, 1)" "tps0 ! j2 = (\hp1\\<^sub>N\<^sub>L, 1)" "tps0 ! j3 = (\N\\<^sub>N, 1)" "tps0 ! (j3 + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j3 + 2) = (\Z\\<^sub>N, 1)" "tps0 ! (j3 + 3) = (\t'\\<^sub>N, 1)" "tps0 ! (j3 + 4) = (\formula_n \\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j3 + 5) = (\formula_n \'\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j3 + 6) = (\t\\<^sub>N, 1)" "\i. 6 < i \ i < 17 \ tps0 ! (j3 + i) = (\[]\, 1)" assumes "tps' = tps0 [1 := nlltape (nss @ formula_n (\ t)), j3 + 6 := (\Suc t\\<^sub>N, 1), j3 + 3 := (\t' - 1\\<^sub>N, 1)]" assumes "ttt = 16114765 * 2^(16*Z) * N ^ 6" shows "transforms (tm_chi j1 j2 j3) tps0 ttt tps'" proof - interpret loc: turing_machine_chi j1 j2 j3 . have G: "H \ 3" using H_gr_2 by simp then have N: "N \ 1" using N_eq by simp have Z: "Z = 3 * H" using Z_def by simp have len_hp0: "length hp0 = Suc T'" using assms by simp have len_hp1: "length hp1 = Suc T'" using assms by simp have hp0: "\i T'" proof standard+ fix i :: nat assume "i < length hp0" then have "hp0 ! i = exc (zeros m) i <#> 0" using map_nth assms(10) by blast then show "hp0 ! i \ T'" using inputpos_less inputpos_def by simp qed have hp1: "\i T'" proof standard+ fix i :: nat assume "i < length hp1" then have "hp1 ! i = exc (zeros m) i <#> 1" using map_nth assms(11) by blast then show "hp1 ! i \ T'" using headpos_1_less by simp qed have psi: "variables \ \ {..<3*Z+H}" "fsize \ \ (3*Z+H) * 2 ^ (3*Z+H)" "length \ \ 2 ^ (3*Z+H)" using psi by simp_all have psi': "variables \' \ {..<2*Z+H}" "fsize \' \ (2*Z+H) * 2 ^ (2*Z+H)" "length \' \ 2 ^ (2*Z+H)" using psi' by simp_all let ?sigma = "[N + (t - 1) * Z.. t then [N + previous hp1 t * Z.. 0" if "i \ T' " for i using that assms map_nth len_hp0 by (metis (no_types, lifting) le_imp_less_Suc) then have hp0_eq_inputpos: "hp0 ! i = inputpos m i" if "i \ T' " for i using inputpos_def that by simp have hp1_nth: "hp1 ! i = exc (zeros m) i <#> 1" if "i \ T' " for i using that assms map_nth len_hp1 by (metis (no_types, lifting) le_imp_less_Suc) have previous_eq_prev: "previous hp1 idx = prev m idx" if "idx \ T' " for idx proof (cases "\ii 1 = exc (zeros m) idx <#> 1" using that hp1_nth by auto then have "prev m idx = (GREATEST i. i < idx \ exc (zeros m) i <#> 1 = exc (zeros m) idx <#> 1)" using prev_def by simp have "previous hp1 idx = (GREATEST i. i < idx \ hp1 ! i = hp1 ! idx)" using True previous_def by simp also have "... = (GREATEST i. i < idx \ exc (zeros m) i <#> 1 = exc (zeros m) idx <#> 1)" (is "Greatest ?P = Greatest ?Q") proof (rule Greatest_equality) have "\i. ?Q i" using 1 by simp moreover have 2: "\y. ?Q y \ y \ idx" by simp ultimately have 3: "?Q (Greatest ?Q)" using GreatestI_ex_nat[of ?Q] by blast then have 4: "Greatest ?Q < idx" by simp then have "Greatest ?Q \ T' " using that by simp then have "hp1 ! (Greatest ?Q) = exc (zeros m) (Greatest ?Q) <#> 1" using hp1_nth by simp moreover have "hp1 ! idx = exc (zeros m) idx <#> 1" using that hp1_nth by simp ultimately have "hp1 ! (Greatest ?Q) = hp1 ! idx" using 3 by simp then show "?P (Greatest ?Q)" using 4 by simp show "i \ Greatest ?Q" if "?P i" for i proof - have "i < idx" using that by simp then have "hp1 ! i = exc (zeros m) i <#> 1" using `idx \ T' ` hp1_nth by simp moreover have "hp1 ! idx = exc (zeros m) idx <#> 1" using `idx \ T' ` hp1_nth by simp ultimately have "exc (zeros m) i <#> 1 = exc (zeros m) idx <#> 1" using that by simp then have "?Q i" using `i < idx` by simp then show ?thesis using Greatest_le_nat[of ?Q i] 2 by blast qed qed also have "... = prev m idx" using prev_def 1 by simp finally show ?thesis . next case False have "\ (\i 1 = exc (zeros m) idx <#> 1)" proof (rule ccontr) assume "\ (\ (\i 1 = exc (zeros m) idx <#> 1))" then obtain i where "i < idx" "exc (zeros m) i <#> 1 = exc (zeros m) idx <#> 1" by auto then have "hp1 ! i = hp1 ! idx" using hp1_nth that by simp then show False using False `i < idx` by simp qed then have "prev m idx = idx" using prev_def by auto moreover have "previous hp1 idx = idx" using False assms previous_def by auto ultimately show ?thesis by simp qed have "\\<^sub>0 i @ \\<^sub>1 i @ \\<^sub>2 i = [N + i * Z..\<^sub>0 i @ \\<^sub>1 i @ \\<^sub>2 i = [N + i * Z.. (inputpos m i) = [inputpos m i * H.. t = ?sigma" if "prev m t < t" proof - have "previous hp1 t \ t" using t that previous_eq_prev by simp then have "?sigma = [N + (t - 1) * Z..\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1)) @ (\\<^sub>0 (prev m t) @ \\<^sub>1 (prev m t) @ \\<^sub>2 (prev m t)) @ (\\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t) @ \ (inputpos m t)" using zeta012 gamma by simp also have "... = \ t" using rho_def by simp finally have "?sigma = \ t" . then show ?thesis by simp qed have rho': "\' t = ?sigma" if "prev m t = t" proof - have "previous hp1 t = t" using t that previous_eq_prev by simp then have "?sigma = [N + (t - 1) * Z..\<^sub>0 (t - 1) @ \\<^sub>1 (t - 1) @ \\<^sub>2 (t - 1)) @ (\\<^sub>0 t @ \\<^sub>1 t @ \\<^sub>2 t) @ \ (inputpos m t)" using zeta012 gamma by simp also have "... = \' t" using rho'_def by simp finally have "?sigma = \' t" . then show ?thesis by simp qed have "\ t = relabel ?sigma (if previous hp1 t = t then \' else \)" proof (cases "prev m t < t") case True then have "\ t = relabel (\ t) \" using chi_def by simp moreover have "previous hp1 t < t" using t True previous_eq_prev by simp ultimately show ?thesis using rho True by simp next case False then have *: "prev m t = t" by (simp add: nat_less_le prev_le) then have "\ t = relabel (\' t) \'" using chi_def by simp moreover have "previous hp1 t = t" using t * previous_eq_prev by simp ultimately show ?thesis using rho' * by simp qed then show "transforms (tm_chi j1 j2 j3) tps0 ttt tps'" using transforms_tm_chi[OF jk G Z N len_hp0 hp0 len_hp1 hp1 t T T'_less psi psi' tps0 _ assms(24)] assms(23) by simp qed lemma Z_sq_le: "Z^2 \ 2^(16*Z)" proof - have "Z^2 \ 2^(2*Z)" using pow2_le_2pow2[of Z] by simp also have "... \ 2^(16*Z)" by simp finally show "Z^2 \ 2^(16*Z)" . qed lemma tm_PHI9 [transforms_intros]: fixes tps0 tps' :: "tape list" and k :: nat and j1 j2 j3 :: tapeidx fixes nss :: "nat list list" assumes jk: "length tps0 = k" "1 < j1" "j1 < j2" "j2 < j3" "j3 + 17 < k" assumes "hp0 = map (\t. exc (zeros m) t <#> 0) [0..t. exc (zeros m) t <#> 1) [0..hp0\\<^sub>N\<^sub>L, 1)" "tps0 ! j2 = (\hp1\\<^sub>N\<^sub>L, 1)" "tps0 ! j3 = (\N\\<^sub>N, 1)" "tps0 ! (j3 + 1) = (\H\\<^sub>N, 1)" "tps0 ! (j3 + 2) = (\Z\\<^sub>N, 1)" "tps0 ! (j3 + 3) = (\T'\\<^sub>N, 1)" "tps0 ! (j3 + 4) = (\formula_n psi\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j3 + 5) = (\formula_n psi'\\<^sub>N\<^sub>L\<^sub>L, 1)" "tps0 ! (j3 + 6) = (\1\\<^sub>N, 1)" "\i. 6 < i \ i < 17 \ tps0 ! (j3 + i) = (\[]\, 1)" assumes tps': "tps' = tps0 [1 := nlltape (nss @ formula_n \\<^sub>9), j3 + 6 := (\Suc T'\\<^sub>N, 1), j3 + 3 := (\0\\<^sub>N, 1)]" assumes "ttt = 16114767 * 2 ^ (16 * Z) * N ^ 7" shows "transforms (tm_PHI9 j1 j2 j3) tps0 ttt tps'" proof - define tps where "tps = (\t. tps0 [1 := nlltape (nss @ concat (map (\t. formula_n (\ (Suc t))) [0..Suc t\\<^sub>N, 1), j3 + 3 := (\T' - t\\<^sub>N, 1)])" have "transforms (tm_PHI9 j1 j2 j3) (tps 0) ttt (tps T')" unfolding tm_PHI9_def proof (tform) let ?ttt = "16114765 * 2^(16*Z) * N ^ 6" show "transforms (tm_chi j1 j2 j3) (tps i) ?ttt (tps (Suc i))" if "i < T' " for i proof (rule tm_chi ; (use assms tps_def that in simp ; fail)?) show "tps (Suc i) = (tps i) [1 := nlltape ((nss @ concat (map (\t. formula_n (\ (Suc t))) [0.. (Suc i))), j3 + 6 := (\Suc (Suc i)\\<^sub>N, 1), j3 + 3 := (\T' - i - 1\\<^sub>N, 1)]" using that tps_def by (simp add: list_update_swap) qed show "\i. i < T' \ read (tps i) ! (j3 + 3) \ \" using jk tps_def read_ncontents_eq_0 by simp show "\ read (tps T') ! (j3 + 3) \ \" using jk tps_def read_ncontents_eq_0 by simp show "T' * (16114765 * 2 ^ (16 * Z) * N ^ 6 + 2) + 1 \ ttt" proof - have "T' * (16114765 * 2 ^ (16 * Z) * N ^ 6 + 2) + 1 \ T' * (16114767 * 2 ^ (16 * Z) * N ^ 6) + 1" using Z_sq_le H_gr_2 N_eq by auto also have "... \ T' * (16114767 * 2 ^ (16 * Z) * N ^ 6) + (16114767 * 2 ^ (16 * Z) * N ^ 6)" using H_gr_2 N_eq by auto also have "... = Suc T' * (16114767 * 2 ^ (16 * Z) * N ^ 6)" by simp also have "... \ N * (16114767 * 2 ^ (16 * Z) * N ^ 6)" using T'_less Suc_leI mult_le_mono1 by presburger also have "... = 16114767 * 2 ^ (16 * Z) * N ^ 7" by algebra also have "... = ttt" using assms(20) by simp finally show ?thesis . qed qed moreover have "tps 0 = tps0" using tps_def tps0 list_update_id[of tps0 "Suc 0"] list_update_id[of tps0 "j3 + 6"] list_update_id[of tps0 "j3 + 3"] by simp moreover have "tps T' = tps'" proof - have "concat (map (\t. formula_n (\ (Suc t))) [0..t. \ (Suc t)) [0..Satisfiability\label{s:Sat}\ theory Satisfiability imports Wellformed NP begin text \ This chapter introduces the language \SAT{} and shows that it is in $\NP$, which constitutes the easier part of the Cook-Levin theorem. The other part, the $\NP$-hardness of \SAT{}, is what all the following chapters are concerned with. We first introduce Boolean formulas in conjunctive normal form and the concept of satisfiability. Then we define a way to represent such formulas as bit strings, leading to the definition of the language \SAT{} as a set of strings (Section~\ref{s:sat-sat}). For the proof that \SAT{} is in $\NP$, we construct a Turing machine that, given a CNF formula and a string representing a variable assignment, outputs \textbf{1} iff. the assignment satisfies the formula. The TM will run in polynomial time, and there are always assignments polynomial (in fact, linear) in the length of the formula (Section~\ref{s:Sat-np}). \ section \The language \SAT{}\label{s:sat-sat}\ text \ \SAT{} is the language of all strings representing satisfiable Boolean formulas in conjunctive normal form (CNF). This section introduces a minimal version of Boolean formulas in conjunctive normal form, including the concepts of assignments and satisfiability. \ subsection \CNF formulas and satisfiability\label{s:CNF}\ text \ Arora and Barak~\cite[p.~44]{ccama} define Boolean formulas in general as expressions over $\land$, $\lor$, $\lnot$, parentheses, and variables $v_1, v_2, \dots$ in the usual way. Boolean formulas in conjunctive normal form are defined as $\bigwedge_i\left(\bigvee_j v_{i_j}\right)$, where the $v_{i_j}$ are literals. This definition does not seem to allow for empty clauses. Also whether the ``empty CNF formula'' exists is somewhat doubtful. Nevertheless, our formalization allows for both empty clauses and the empty CNF formula, because this enables us to represent CNF formulas as lists of clauses and clauses as lists of literals without having to somehow forbid the empty list. This seems to be a popular approach for formalizing CNF formulas in the context of \SAT{} and $\NP$~\cite{Gher2021MechanisingCT,Multiset_Ordering_NPC-AFP}. We identify a variable $v_i$ with its index $i$, which can be any natural number. A \emph{literal} can either be positive or negative, representing a variable or negated variable, respectively. \null \ datatype literal = Neg nat | Pos nat type_synonym clause = "literal list" type_synonym formula = "clause list" text \ An \emph{assignment} maps all variables, given by their index, to a Boolean: \ type_synonym assignment = "nat \ bool" abbreviation satisfies_literal :: "assignment \ literal \ bool" where "satisfies_literal \ x \ case x of Neg n \ \ \ n | Pos n \ \ n" definition satisfies_clause :: "assignment \ clause \ bool" where "satisfies_clause \ c \ \x\set c. satisfies_literal \ x" definition satisfies :: "assignment \ formula \ bool" (infix "\" 60) where "\ \ \ \ \c\set \. satisfies_clause \ c" text \ As is customary, the empty clause is satisfied by no assignment, and the empty CNF formula is satisfied by every assignment. \ proposition "\ satisfies_clause \ []" by (simp add: satisfies_clause_def) proposition "\ \ []" by (simp add: satisfies_def) lemma satisfies_clause_take: assumes "i < length clause" shows "satisfies_clause \ (take (Suc i) clause) \ satisfies_clause \ (take i clause) \ satisfies_literal \ (clause ! i)" using assms satisfies_clause_def by (auto simp add: take_Suc_conv_app_nth) lemma satisfies_take: assumes "i < length \" shows "\ \ take (Suc i) \ \ \ \ take i \ \ satisfies_clause \ (\ ! i)" using assms satisfies_def by (auto simp add: take_Suc_conv_app_nth) lemma satisfies_append: assumes "\ \ \\<^sub>1 @ \\<^sub>2" shows "\ \ \\<^sub>1" and "\ \ \\<^sub>2" using assms satisfies_def by simp_all lemma satisfies_append': assumes "\ \ \\<^sub>1" and "\ \ \\<^sub>2" shows "\ \ \\<^sub>1 @ \\<^sub>2" using assms satisfies_def by auto lemma satisfies_concat_map: assumes "\ \ concat (map f [0.. \ f i" using assms satisfies_def by simp lemma satisfies_concat_map': assumes "\i. i < k \ \ \ f i" shows "\ \ concat (map f [0.. The main ingredient for defining \SAT{} is the concept of \emph{satisfiable} CNF formula: \ definition satisfiable :: "formula \ bool" where "satisfiable \ \ \\. \ \ \" text \ The set of all variables used in a CNF formula is finite. \ definition variables :: "formula \ nat set" where "variables \ \ {n. \c\set \. Neg n \ set c \ Pos n \ set c}" lemma finite_variables: "finite (variables \)" proof - define voc :: "clause \ nat set" where "voc c = {n. Neg n \ set c \ Pos n \ set c}" for c let ?vocs = "set (map voc \)" have "finite (voc c)" for c proof (induction c) case Nil then show ?case using voc_def by simp next case (Cons a c) have "voc (a # c) = {n. Neg n \ set (a # c) \ Pos n \ set (a # c)}" using voc_def by simp also have "... = {n. Neg n \ set c \ Neg n = a \ Pos n \ set c \ Pos n = a}" by auto also have "... = {n. (Neg n \ set c \ Pos n \ set c) \ (Pos n = a \ Neg n = a)}" by auto also have "... = {n. Neg n \ set c \ Pos n \ set c} \ {n. Pos n = a \ Neg n = a}" by auto also have "... = voc c \ {n. Pos n = a \ Neg n = a}" using voc_def by simp finally have "voc (a # c) = voc c \ {n. Pos n = a \ Neg n = a}" . moreover have "finite {n. Pos n = a \ Neg n = a}" using finite_nat_set_iff_bounded by auto ultimately show ?case using Cons by simp qed moreover have "variables \ = \?vocs" using variables_def voc_def by auto moreover have "finite ?vocs" by simp ultimately show ?thesis by simp qed lemma variables_append: "variables (\\<^sub>1 @ \\<^sub>2) = variables \\<^sub>1 \ variables \\<^sub>2" using variables_def by auto text \ Arora and Barak~\cite[Claim~2.13]{ccama} define the \emph{size} of a CNF formula as the numbr of $\wedge / \vee$ symbols. We use a slightly different definition, namely the number of literals: \ definition fsize :: "formula \ nat" where "fsize \ \ sum_list (map length \)" subsection \Predicates on assignments\ text \ Every CNF formula is satisfied by a set of assignments. Conversely, for certain sets of assignments we can construct CNF formulas satisfied by exactly these assignments. This will be helpful later when we construct formulas for reducing arbitrary languages to \SAT{} (see Section~\ref{s:Reducing}). \ subsubsection \Universality of CNF formulas\ text \ A set (represented by a predicate) $F$ of assignments depends on the first $\ell$ variables iff.\ any two assignments that agree on the first $\ell$ variables are either both in the set or both outside of the set. \ definition depon :: "nat \ (assignment \ bool) \ bool" where "depon l F \ \\\<^sub>1 \\<^sub>2. (\i\<^sub>1 i = \\<^sub>2 i) \ F \\<^sub>1 = F \\<^sub>2" text \ Lists of all strings of the same length: \ fun str_of_len :: "nat \ string list" where "str_of_len 0 = [[]]" | "str_of_len (Suc l) = map ((#) \) (str_of_len l) @ map ((#) \) (str_of_len l)" lemma length_str_of_len: "length (str_of_len l) = 2 ^ l" by (induction l) simp_all lemma in_str_of_len_length: "xs \ set (str_of_len l) \ length xs = l" by (induction l arbitrary: xs) auto lemma length_in_str_of_len: "length xs = l \ xs \ set (str_of_len l)" proof (induction l arbitrary: xs) case 0 then show ?case by simp next case (Suc l) then obtain y ys where "xs = y # ys" by (meson length_Suc_conv) then have "length ys = l" using Suc by simp show ?case proof (cases y) case True then have "xs \ set (map ((#) \) (str_of_len l))" using `length ys = l` Suc `xs = y # ys` by simp then show ?thesis by simp next case False then have "xs \ set (map ((#) \) (str_of_len l))" using `length ys = l` Suc `xs = y # ys` by simp then show ?thesis by simp qed qed text \ A predicate $F$ depending on the first $\ell$ variables $v_0, \dots, v_{\ell-1}$ can be regarded as a truth table over $\ell$ variables. The next lemma shows that for every such truth table there exists a CNF formula with at most $2^\ell$ clauses and $\ell\cdot2^\ell$ literals over the first $\ell$ variables. This is the well-known fact that every Boolean function (over $\ell$ variables) can be represented by a CNF formula~\cite[Claim~2.13]{ccama}. \ lemma depon_ex_formula: assumes "depon l F" shows "\\. fsize \ \ l * 2 ^ l \ length \ \ 2 ^ l \ variables \ \ {.. (\\. F \ = \ \ \)" proof - define cl where "cl = (\v. map (\i. if v ! i then Neg i else Pos i) [0.. map a [0.. v ! i" using \length v = l\ \v \ map a [0.. by (smt (verit, best) atLeastLessThan_iff map_eq_conv map_nth set_upt) then have *: "cl v ! i = (if v ! i then Neg i else Pos i)" using cl_def by simp then have "case (cl v ! i) of Neg n \ \ a n | Pos n \ a n" using i(2) by simp then show ?thesis using cl_def * that(1) satisfies_clause_def i(1) by fastforce qed have cl2: "v \ map a [0..x\set (cl v). case x of Neg n \ \ a n | Pos n \ a n" using satisfies_clause_def by simp then obtain i where i: "i < l" and "case (cl v ! i) of Neg n \ \ a n | Pos n \ a n" using cl_def by auto then have "v ! i \ a i" using cl_def by fastforce then show False using i assm by simp qed have filter_length_nth: "f (vs ! j)" if "vs = filter f sol" and "j < length vs" for vs sol :: "'a list" and f j using that nth_mem by (metis length_removeAll_less less_irrefl removeAll_filter_not) have sum_list_map: "sum_list (map g xs) \ k * length xs" if "\x. x \ set xs \ g x = k" for xs :: "'a list" and g k using that proof (induction "length xs" arbitrary: xs) case 0 then show ?case by simp next case (Suc x) then obtain y ys where "xs = y # ys" by (metis length_Suc_conv) then have "length ys = x" using Suc by simp have "y \ set xs" using `xs = y # ys` by simp have "sum_list (map g xs) = sum_list (map g (y # ys))" using `xs = y # ys` by simp also have "... = g y + sum_list (map g ys)" by simp also have "... = k + sum_list (map g ys)" using Suc `y \ set xs` by simp also have "... \ k + k * length ys" using Suc `length ys = x` \xs = y # ys\ by auto also have "... = k * length xs" by (metis Suc.hyps(2) \length ys = x\ mult_Suc_right) finally show ?case by simp qed define vs where "vs = filter (\v. F (\i. if i < l then v ! i else False) = False) (str_of_len l)" define \ where "\ = map cl vs" have "a \ \" if "F a" for a proof - define v where "v = map a [0..i. if i < l then v ! i else False) j = a j" if "j < l" for j by (simp add: that) then have *: "F (\i. if i < l then v ! i else False)" using assms(1) depon_def that by (smt (verit, ccfv_SIG)) have "satisfies_clause a c" if "c \ set \" for c proof - obtain j where j: "c = \ ! j" "j < length \" using \_def `c \ set \` by (metis in_set_conv_nth) then have "c = cl (vs ! j)" using \_def by simp have "j < length vs" using \_def j by simp then have "F (\i. if i < l then (vs ! j) ! i else False) = False" using vs_def filter_length_nth by blast then have "vs ! j \ v" using * by auto moreover have "length (vs ! j) = l" using vs_def length_str_of_len \j < length vs\ - by (smt (z3) filter_eq_nths in_str_of_len_length notin_set_nthsI nth_mem) + by (smt (verit) filter_eq_nths in_str_of_len_length notin_set_nthsI nth_mem) ultimately have "satisfies_clause a (cl (vs ! j))" using v_def cl1 by simp then show ?thesis using `c = cl (vs ! j)` by simp qed then show ?thesis using satisfies_def by simp qed moreover have "F \" if "\ \ \" for \ proof (rule ccontr) assume assm: "\ F \" define v where "v = map \ [0..i. if i < l then v ! i else False) = False" proof - have "(\i. if i < l then v ! i else False) j = \ j" if "j < l" for j using v_def by (simp add: that) then show ?thesis using assm assms(1) depon_def by (smt (verit, best)) qed have "length v = l" using v_def by simp then obtain j where "j < length (str_of_len l)" and "v = str_of_len l ! j" by (metis in_set_conv_nth length_in_str_of_len) then have "v \ set vs" using vs_def * by fastforce then have "cl v \ set \" using \_def by simp then have "satisfies_clause \ (cl v)" using that satisfies_def by simp then have "v \ map \ [0..\. F \ = \ \ \" by auto moreover have "fsize \ \ l * 2 ^ l" proof - have "length c = l" if "c \ set \" for c using that cl_def \_def by auto then have "fsize \ \ l * length \" unfolding fsize_def using sum_list_map by auto also have "... \ l * length (str_of_len l)" using \_def vs_def by simp also have "... = l * 2 ^ l" using length_str_of_len by simp finally show ?thesis . qed moreover have "length \ \ 2 ^ l" proof - have "length \ \ length (str_of_len l)" using \_def vs_def by simp also have "... = 2 ^ l" using length_str_of_len by simp finally show ?thesis . qed moreover have "variables \ \ {.. variables \" then obtain c where c: "c \ set \" "Neg x \ set c \ Pos x \ set c" using variables_def by auto then obtain v where v: "v \ set (str_of_len l)" "c = cl v" using \_def vs_def by auto then show "x \ {..Substitutions of variables\ text \ We will sometimes consider CNF formulas over the first $\ell$ variables and derive other CNF formulas from them by substituting these variables. Such a substitution will be represented by a list $\sigma$ of length at least $\ell$, meaning that the variable $v_i$ is replaced by $v_{\sigma(i)}$. We will call this operation on formulas \emph{relabel}, and the corresponding one on literals \emph{rename}: \ fun rename :: "nat list \ literal \ literal" where "rename \ (Neg i) = Neg (\ ! i)" | "rename \ (Pos i) = Pos (\ ! i)" definition relabel :: "nat list \ formula \ formula" where "relabel \ \ \ map (map (rename \)) \" lemma fsize_relabel: "fsize (relabel \ \) = fsize \" using relabel_def fsize_def by (metis length_concat length_map map_concat) text \ A substitution $\sigma$ can also be applied to an assignment and to a list of variable indices: \ definition remap :: "nat list \ assignment \ assignment" where "remap \ \ i \ if i < length \ then \ (\ ! i) else \ i" definition reseq :: "nat list \ nat list \ nat list" where "reseq \ vs \ map ((!) \) vs" lemma length_reseq [simp]: "length (reseq \ vs) = length vs" using reseq_def by simp text \ Relabeling a formula and remapping an assignment are equivalent in a sense. \ lemma satisfies_sigma: assumes "variables \ \ {..}" shows "\ \ relabel \ \ \ remap \ \ \ \" proof assume sat: "\ \ relabel \ \" have "satisfies_clause (remap \ \) c" if "c \ set \" for c proof - obtain i where "i < length \" "\ ! i = c" by (meson \c \ set \\ in_set_conv_nth) then have "satisfies_clause \ (map (rename \) c)" (is "satisfies_clause \ ?c") using sat satisfies_def relabel_def by auto then obtain x where "x\set ?c" "case x of Neg n \ \ \ n | Pos n \ \ n" using satisfies_clause_def by auto then obtain j where j: "j < length ?c" "case (?c ! j) of Neg n \ \ \ n | Pos n \ \ n" by (metis in_set_conv_nth) have "case c ! j of Neg n \ \ (remap \ \) n | Pos n \ (remap \ \) n" proof (cases "c ! j") case (Neg n) then have 1: "?c ! j = Neg (\ ! n)" using j(1) by simp have "n \ variables \" using Neg j(1) nth_mem that variables_def by force then have "n < length \" using assms by auto then show ?thesis using Neg 1 j(2) remap_def by auto next case (Pos n) then have 1: "?c ! j = Pos (\ ! n)" using j(1) by simp have "n \ variables \" using Pos j(1) nth_mem that variables_def by force then have "n < length \" using assms by auto then show ?thesis using Pos 1 j(2) remap_def by auto qed then show ?thesis using satisfies_clause_def j by auto qed then show "remap \ \ \ \" using satisfies_def by simp next assume sat: "remap \ \ \ \" have "satisfies_clause \ c" if "c \ set (relabel \ \)" for c proof - let ?phi = "relabel \ \" let ?beta = "remap \ \" obtain i where i: "i < length ?phi" "?phi ! i = c" by (meson \c \ set ?phi\ in_set_conv_nth) then have "satisfies_clause ?beta (\ ! i)" (is "satisfies_clause _ ?c") using sat satisfies_def relabel_def by simp then obtain x where "x\set ?c" "case x of Neg n \ \ ?beta n | Pos n \ ?beta n" using satisfies_clause_def by auto then obtain j where j: "j < length ?c" "case (?c ! j) of Neg n \ \ ?beta n | Pos n \ ?beta n" by (metis in_set_conv_nth) then have ren: "c ! j = rename \ (?c ! j)" using i relabel_def by auto have "case c ! j of Neg n \ \ \ n | Pos n \ \ n" proof (cases "?c ! j") case (Neg n) then have *: "c ! j = Neg (\ ! n)" by (simp add: ren) have "n \ variables \" using Neg i j variables_def that length_map mem_Collect_eq nth_mem relabel_def by force then have "n < length \" using assms by auto moreover have "\ (remap \ \) n" using j(2) Neg by simp ultimately have "\ \ (\ ! n)" using remap_def by simp then show ?thesis by (simp add: *) next case (Pos n) then have *: "c ! j = Pos (\ ! n)" by (simp add: ren) have "n \ variables \" using Pos i j variables_def that length_map mem_Collect_eq nth_mem relabel_def by force then have "n < length \" using assms by auto moreover have "(remap \ \) n" using j(2) Pos by simp ultimately have "\ (\ ! n)" using remap_def by simp then show ?thesis by (simp add: *) qed moreover have "length c = length (\ ! i)" using relabel_def i by auto ultimately show ?thesis using satisfies_clause_def j by auto qed then show "\ \ relabel \ \" by (simp add: satisfies_def) qed subsection \Representing CNF formulas as strings\label{s:sat-sat-repr}\ text \ By identifying negated literals with even numbers and positive literals with odd numbers, we can identify literals with natural numbers. This yields a straightforward representation of a clause as a list of numbers and of a CNF formula as a list of lists of numbers. Such a list can, in turn, be represented as a symbol sequence over a quaternary alphabet as described in Section~\ref{s:tm-numlistlist}, which ultimately can be encoded over a binary alphabet (see Section~\ref{s:tm-quaternary}). This is essentially how we represent CNF formulas as strings. We have to introduce a bunch of functions for mapping between these representations. \null \ fun literal_n :: "literal \ nat" where "literal_n (Neg i) = 2 * i" | "literal_n (Pos i) = Suc (2 * i)" definition n_literal :: "nat \ literal" where "n_literal n \ if even n then Neg (n div 2) else Pos (n div 2)" lemma n_literal_n: "n_literal (literal_n x) = x" using n_literal_def by (cases x) simp_all lemma literal_n_literal: "literal_n (n_literal n) = n" using n_literal_def by simp definition clause_n :: "clause \ nat list" where "clause_n cl \ map literal_n cl" definition n_clause :: "nat list \ clause" where "n_clause ns \ map n_literal ns" lemma n_clause_n: "n_clause (clause_n cl) = cl" using n_clause_def clause_n_def n_literal_n by (simp add: map_idI) lemma clause_n_clause: "clause_n (n_clause n) = n" using n_clause_def clause_n_def literal_n_literal by (simp add: map_idI) definition formula_n :: "formula \ nat list list" where "formula_n \ \ map clause_n \" definition n_formula :: "nat list list \ formula" where "n_formula nss \ map n_clause nss" lemma n_formula_n: "n_formula (formula_n \) = \" using n_formula_def formula_n_def n_clause_n by (simp add: map_idI) lemma formula_n_formula: "formula_n (n_formula nss) = nss" using n_formula_def formula_n_def clause_n_clause by (simp add: map_idI) definition formula_zs :: "formula \ symbol list" where "formula_zs \ \ numlistlist (formula_n \)" text \ The mapping between formulas and well-formed symbol sequences for lists of lists of numbers is bijective. \ lemma formula_n_inj: "formula_n \\<^sub>1 = formula_n \\<^sub>2 \ \\<^sub>1 = \\<^sub>2" using n_formula_n by metis definition zs_formula :: "symbol list \ formula" where "zs_formula zs \ THE \. formula_zs \ = zs" lemma zs_formula: assumes "numlistlist_wf zs" shows "\!\. formula_zs \ = zs" proof - obtain nss where nss: "numlistlist nss = zs" using assms numlistlist_wf_def by auto let ?phi = "n_formula nss" have *: "formula_n ?phi = nss" using nss formula_n_formula by simp then have "formula_zs ?phi = zs" using nss formula_zs_def by simp then have "\\. formula_zs \ = zs" by auto moreover have "\ = ?phi" if "formula_zs \ = zs" for \ proof - have "numlistlist (formula_n \) = zs" using that formula_zs_def by simp then have "nss = formula_n \" using nss numlistlist_inj by simp then show ?thesis using formula_n_inj * by simp qed ultimately show ?thesis by auto qed lemma zs_formula_zs: "zs_formula (formula_zs \) = \" by (simp add: formula_n_inj formula_zs_def numlistlist_inj the_equality zs_formula_def) lemma formula_zs_formula: assumes "numlistlist_wf zs" shows "formula_zs (zs_formula zs) = zs" using assms zs_formula zs_formula_zs by metis text \ There will of course be Turing machines that perform computations on formulas. In order to bound their running time, we need bounds for the length of the symbol representation of formulas. \ lemma nlength_literal_n_Pos: "nlength (literal_n (Pos n)) \ Suc (nlength n)" using nlength_times2plus1 by simp lemma nlength_literal_n_Neg: "nlength (literal_n (Neg n)) \ Suc (nlength n)" using nlength_times2 by simp lemma nlllength_formula_n: fixes V :: nat and \ :: formula assumes "\v. v \ variables \ \ v \ V" shows "nlllength (formula_n \) \ fsize \ * Suc (Suc (nlength V)) + length \" using assms proof (induction \) case Nil then show ?case using formula_n_def by simp next case (Cons cl \) then have 0: "\v. v \ variables \ \ v \ V" using variables_def by simp have 1: "n \ V" if "Pos n \ set cl" for n using that variables_def by (simp add: Cons.prems) have 2: "n \ V" if "Neg n \ set cl" for n using that variables_def by (simp add: Cons.prems) have 3: "nlength (literal_n v) \ Suc (nlength V)" if "v \ set cl" for v proof (cases v) case (Neg n) then have "nlength (literal_n v) \ Suc (nlength n)" using nlength_literal_n_Neg by blast moreover have "n \ V" using Neg that 2 by simp ultimately show ?thesis using nlength_mono by fastforce next case (Pos n) then have "nlength (literal_n v) \ Suc (nlength n)" using nlength_literal_n_Pos by blast moreover have "n \ V" using Pos that 1 by simp ultimately show ?thesis using nlength_mono by fastforce qed have "nllength (clause_n cl) = length (numlist (map literal_n cl))" using clause_n_def nllength_def by simp have "nllength (clause_n cl) = (\n\(map literal_n cl). Suc (nlength n))" using clause_n_def nllength by simp also have "... = (\v\cl. Suc (nlength (literal_n v)))" proof - have "map (\n. Suc (nlength n)) (map literal_n cl) = map (\v. Suc (nlength (literal_n v))) cl" by simp then show ?thesis by metis qed also have "... \ (\v\cl. Suc (Suc (nlength V)))" using sum_list_mono[of cl "\v. Suc (nlength (literal_n v))" "\v. Suc (Suc (nlength V))"] 3 by simp also have "... = Suc (Suc (nlength V)) * length cl" using sum_list_const by blast finally have 4: "nllength (clause_n cl) \ Suc (Suc (nlength V)) * length cl" . have "concat (map (\ns. numlist ns @ [5]) (map clause_n (cl # \))) = (numlist (clause_n cl) @ [5]) @ concat (map (\ns. numlist ns @ [5]) (map clause_n \))" by simp then have "length (concat (map (\ns. numlist ns @ [5]) (map clause_n (cl # \)))) = length ((numlist (clause_n cl) @ [5]) @ concat (map (\ns. numlist ns @ [5]) (map clause_n \)))" by simp then have "nlllength (formula_n (cl # \)) = length ((numlist (clause_n cl) @ [5]) @ concat (map (\ns. numlist ns @ [5]) (map clause_n \)))" using formula_n_def numlistlist_def nlllength_def by simp also have "... = length (numlist (clause_n cl) @ [5]) + length (concat (map (\ns. numlist ns @ [5]) (map clause_n \)))" by simp also have "... = length (numlist (clause_n cl) @ [5]) + nlllength (formula_n \)" using formula_n_def numlistlist_def nlllength_def by simp also have "... = Suc (nllength (clause_n cl)) + nlllength (formula_n \)" using nllength_def by simp also have "... \ Suc (Suc (Suc (nlength V)) * length cl) + nlllength (formula_n \)" using 4 by simp also have "... \ Suc (Suc (Suc (nlength V)) * length cl) + fsize \ * Suc (Suc (nlength V)) + length \" using Cons 0 by simp also have "... = fsize (cl # \) * Suc (Suc (nlength V)) + length (cl # \)" by (simp add: add_mult_distrib2 mult.commute fsize_def) finally show ?case by simp qed text \ Since \SAT{} is supposed to be a set of strings rather than symbol sequences, we need to map symbol sequences representing formulas to strings: \ abbreviation formula_to_string :: "formula \ string" where "formula_to_string \ \ symbols_to_string (binencode (numlistlist (formula_n \)))" lemma formula_to_string_inj: assumes "formula_to_string \\<^sub>1 = formula_to_string \\<^sub>2" shows "\\<^sub>1 = \\<^sub>2" proof - let ?xs1 = "binencode (numlistlist (formula_n \\<^sub>1))" let ?xs2 = "binencode (numlistlist (formula_n \\<^sub>2))" have bin1: "binencodable (numlistlist (formula_n \\<^sub>1))" by (simp add: Suc_le_eq numeral_2_eq_2 proper_symbols_numlistlist symbols_lt_numlistlist) then have "bit_symbols ?xs1" using bit_symbols_binencode by simp then have 1: "string_to_symbols (symbols_to_string ?xs1) = ?xs1" using bit_symbols_to_symbols by simp have bin2: "binencodable (numlistlist (formula_n \\<^sub>2))" by (simp add: Suc_le_eq numeral_2_eq_2 proper_symbols_numlistlist symbols_lt_numlistlist) then have "bit_symbols ?xs2" using bit_symbols_binencode by simp then have "string_to_symbols (symbols_to_string ?xs2) = ?xs2" using bit_symbols_to_symbols by simp then have "?xs1 = ?xs2" using 1 assms by simp then have "numlistlist (formula_n \\<^sub>1) = numlistlist (formula_n \\<^sub>2)" using binencode_inj bin1 bin2 by simp then have "formula_n \\<^sub>1 = formula_n \\<^sub>2" using numlistlist_inj by simp then show ?thesis using formula_n_inj by simp qed text \ While @{const formula_to_string} maps every CNF formula to a string, not every string represents a CNF formula. We could just ignore such invalid strings and define \SAT{} to only contain well-formed strings. But this would implicitly place these invalid strings in the complement of \SAT{}. While this does not cause us any issues here, it would if we were to introduce co-$\NP$ and wanted to show that $\overline{\mathtt{SAT}}$ is in co-$\NP$, as we would then have to deal with the invalid strings. So it feels a little like cheating to ignore the invalid strings like this. Arora and Barak~\cite[p.~45 footnote~3]{ccama} recommend mapping invalid strings to ``some fixed formula''. A natural candidate for this fixed formula is the empty CNF, since an invalid string in a sense represents nothing, and the empty CNF formula is represented by the empty string. Since the empty CNF formula is satisfiable this implies that all invalid strings become elements of \SAT{}. We end up with the following definition of the protagonist of this article: \ definition SAT :: language where "SAT \ {formula_to_string \ | \. satisfiable \} \ {x. \ (\\. x = formula_to_string \)}" section \\SAT{} is in $\NP$\label{s:Sat-np}\ text \ In order to show that \SAT{} is in $\NP$, we will construct a polynomial-time Turing machine $M$ and specify a polynomial function $p$ such that for every $x$, $x\in \SAT$ iff. there is a $u\in\bbOI^{p(|x|)}$ such that $M$ outputs \textbf{1} on $\langle x; u\rangle$. The idea is straightforward: Let $\phi$ be the formula represented by the string $x$. Interpret the string $u$ as a list of variables and interpret this as the assignment that assigns True to only the variables in the list. Then check if the assignment satisfies the formula. This check is ``obviously'' possible in polynomial time because $M$ simply has to iterate over all clauses and check if at least one literal per clause is true under the assignment. Checking if a literal is true is simply a matter of checking whether the literal's variable is in the list $u$. If the assignment satisfies $\phi$, output \textbf{1}, otherwise the empty symbol sequence. If $\phi$ is unsatisfiable then no assignment, hence no $u$ no matter the length will make $M$ output \textbf{1}. On the other hand, if $\phi$ is satisfiable then there will be a satisfying assignment where a subset of the variables in $\phi$ are assigned True. Hence there will be a list of variables of at most roughly the length of $\phi$. So setting the polynomial $p$ to something like $p(n) = n$ should suffice. In fact, as we shall see, $p(n) = n$ suffices. This is so because in our representation, the string $x$, being a list of lists, has slightly more overhead per number than the plain list $u$ has. Therefore listing all variables in $\phi$ is guaranteed to need fewer symbols than $x$ has. There are several technical details to work out. First of all, the input to $M$ need not be a well-formed pair. And if it is, the pair $\langle x, u\rangle$ has to be decoded into separate components $x$ and $u$. These have to be decoded again from the binary to the quaternary alphabet, which is only possible if both $x$ and $u$ comprise only bit symbols (\textbf{01}). Then $M$ needs to check if the decoded $x$ and $u$ are valid symbol sequences for lists (of lists) of numbers. In the case of $u$ this is particularly finicky because the definition of $\NP$ requires us to provide a string $u$ of exactly the length $p(|x|)$ and so we have to pad $u$ with extra symbols, which have to be stripped again before the validation can take place. In the first subsection we describe what the verifier TM has to do in terms of symbol sequences. In the subsections after that we devise a Turing machine that implements this behavior. \ subsection \Verifying \SAT{} instances\ text \ Our verifier Turing machine for \SAT{} will implement the following function; that is, on input @{term zs} it will output @{term "verify_sat zs"}. It performs a number of decodings and well-formedness checks and outputs either \textbf{1} or the empty symbol sequence. \ definition verify_sat :: "symbol list \ symbol list" where "verify_sat zs \ let ys = bindecode zs; xs = bindecode (first ys); vs = rstrip \ (bindecode (second ys)) in if even (length (first ys)) \ bit_symbols (first ys) \ numlistlist_wf xs then if bit_symbols (second ys) \ numlist_wf vs then if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then [\] else [] else [] else [\]" text \ Next we show that @{const verify_sat} behaves as is required of a verifier TM for \SAT. Its polynomial running time is the subject of the next subsection. \ text \ First we consider the case that @{term zs} encodes a pair $\langle x, u\rangle$ of strings where $x$ does not represent a CNF formula. Such an $x$ is in \SAT{}, hence the verifier TM is supposed to output \textbf{1}. \ lemma ex_phi_x: assumes "xs = string_to_symbols x" assumes "even (length xs)" and "numlistlist_wf (bindecode xs)" shows "\\. x = formula_to_string \" proof - obtain nss where "numlistlist nss = bindecode xs" using assms(3) numlistlist_wf_def by auto moreover have "binencode (bindecode xs) = xs" using assms(1,2) binencode_decode by simp ultimately have "binencode (numlistlist nss) = xs" by simp then have "binencode (numlistlist (formula_n (n_formula nss))) = xs" using formula_n_formula by simp then have "formula_to_string (n_formula nss) = symbols_to_string xs" by simp then show ?thesis using assms(1) symbols_to_string_to_symbols by auto qed lemma verify_sat_not_wf_phi: assumes "zs = \x; u\" and "\ (\\. x = formula_to_string \)" shows "verify_sat zs = [\]" proof - define ys where "ys = bindecode zs" then have first_ys: "first ys = string_to_symbols x" using first_pair assms(1) by simp then have 2: "bit_symbols (first ys)" using assms(1) bit_symbols_first ys_def by presburger define xs where "xs = bindecode (first ys)" then have "\ (even (length (first ys)) \ bit_symbols (first ys) \ numlistlist_wf xs)" using first_ys ex_phi_x assms(2) by auto then show "verify_sat zs = [\]" unfolding verify_sat_def Let_def using ys_def xs_def by simp qed text \ The next case is that @{term zs} represents a pair $\langle x, u\rangle$ where $x$ represents an unsatisfiable CNF formula. This $x$ is thus not in \SAT{} and the verifier TM must output something different from \textbf{1}, such as the empty string, regardless of $u$. \ lemma verify_sat_not_sat: fixes \ :: formula assumes "zs = \formula_to_string \; u\" and "\ satisfiable \" shows "verify_sat zs = []" proof - have bs_phi: "bit_symbols (binencode (formula_zs \))" using bit_symbols_binencode formula_zs_def proper_symbols_numlistlist symbols_lt_numlistlist by (metis Suc_le_eq dual_order.refl numeral_2_eq_2) define ys where "ys = bindecode zs" then have "first ys = string_to_symbols (formula_to_string \)" using first_pair assms(1) by simp then have first_ys: "first ys = binencode (formula_zs \)" using bit_symbols_to_symbols bs_phi formula_zs_def by simp then have 2: "bit_symbols (first ys)" using assms(1) bit_symbols_first ys_def by presburger have 22: "even (length (first ys))" using first_ys by simp define xs where "xs = bindecode (first ys)" define vs where "vs = rstrip 5 (bindecode (second ys))" have wf_xs: "numlistlist_wf xs" using xs_def first_ys bindecode_encode formula_zs_def numlistlist_wf_def numlistlist_wf_has2' by (metis le_simps(3) numerals(2)) have \: "zs_formula xs = \" using xs_def first_ys "2" binencode_decode formula_to_string_inj formula_zs_def formula_zs_formula wf_xs by auto have "verify_sat zs = (if bit_symbols (second ys) \ numlist_wf vs then if (\v. v \ set (zs_numlist vs)) \ \ then [3] else [] else [])" unfolding verify_sat_def Let_def using ys_def xs_def vs_def 2 22 wf_xs \ by simp then show "verify_sat zs = []" using assms(2) satisfiable_def by simp qed text \ Next we consider the case in which @{term zs} represents a pair $\langle x, u\rangle$ where $x$ represents a satisfiable CNF formula and $u$ a list of numbers padded at the right with @{text \} symbols. This $u$ thus represents a variable assignment, namely the one assigning True to exactly the variables in the list. The verifier TM is to output \textbf{1} iff.\ this assignment satisfies the CNF formula represented by $x$. First we show that stripping away @{text \} symbols does not damage a symbol sequence representing a list of numbers. \ lemma rstrip_numlist_append: "rstrip \ (numlist vars @ replicate pad \) = numlist vars" (is "rstrip \ ?zs = ?ys") proof - have "(LEAST i. i \ length ?zs \ set (drop i ?zs) \ {\}) = length ?ys" proof (rule Least_equality) show "length ?ys \ length ?zs \ set (drop (length ?ys) ?zs) \ {\}" by auto show "length ?ys \ m" if "m \ length ?zs \ set (drop m ?zs) \ {\}" for m proof (rule ccontr) assume "\ length ?ys \ m" then have "m < length ?ys" by simp then have "?ys ! m \ set (drop m ?ys)" by (metis Cons_nth_drop_Suc list.set_intros(1)) moreover have "set (drop m ?ys) \ {\}" using that by simp ultimately have "?ys ! m = \" by auto moreover have "?ys ! m < \" using symbols_lt_numlist `m < length ?ys` by simp ultimately show False by simp qed qed then show ?thesis using rstrip_def by simp qed lemma verify_sat_wf: fixes \ :: formula and pad :: nat assumes "zs = \formula_to_string \; symbols_to_string (binencode (numlist vars @ replicate pad \))\" shows "verify_sat zs = (if (\v. v \ set vars) \ \ then [\] else [])" proof - have bs_phi: "bit_symbols (binencode (formula_zs \))" using bit_symbols_binencode formula_zs_def proper_symbols_numlistlist symbols_lt_numlistlist by (metis Suc_le_eq dual_order.refl numeral_2_eq_2) have "binencodable (numlist vars @ replicate pad \)" using proper_symbols_numlist symbols_lt_numlist binencodable_append[of "numlist vars" "replicate pad \"] by fastforce then have bs_vars: "bit_symbols (binencode (numlist vars @ replicate pad \))" using bit_symbols_binencode by simp define ys where "ys = bindecode zs" then have "first ys = string_to_symbols (formula_to_string \)" using first_pair assms(1) by simp then have first_ys: "first ys = binencode (formula_zs \)" using bit_symbols_to_symbols bs_phi formula_zs_def by simp then have bs_first: "bit_symbols (first ys)" using assms(1) bit_symbols_first ys_def by presburger have even: "even (length (first ys))" using first_ys by simp have second_ys: "second ys = binencode (numlist vars @ replicate pad \)" using bs_vars ys_def assms(1) bit_symbols_to_symbols second_pair by simp then have bs_second: "bit_symbols (second ys)" using bs_vars by simp define xs where "xs = bindecode (first ys)" define vs where "vs = rstrip \ (bindecode (second ys))" then have "vs = rstrip \ (numlist vars @ replicate pad \)" using second_ys \binencodable (numlist vars @ replicate pad \)\ bindecode_encode by simp then have vs: "vs = numlist vars" using rstrip_numlist_append by simp have wf_xs: "numlistlist_wf xs" using xs_def first_ys bindecode_encode formula_zs_def numlistlist_wf_def numlistlist_wf_has2' by (metis le_simps(3) numerals(2)) have "verify_sat zs = (if even (length (first ys)) \ bit_symbols (first ys) \ numlistlist_wf xs then if bit_symbols (second ys) \ numlist_wf vs then if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then [\] else [] else [] else [3])" unfolding verify_sat_def Let_def using bs_second ys_def xs_def vs_def by simp then have *: "verify_sat zs = (if bit_symbols (second ys) \ numlist_wf vs then if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then [\] else [] else [])" unfolding verify_sat_def Let_def using even bs_first wf_xs by simp have "xs = formula_zs \" using xs_def bindecode_encode formula_zs_def first_ys proper_symbols_numlistlist symbols_lt_numlistlist by (simp add: Suc_leI numerals(2)) then have \: "\ = zs_formula xs" by (simp add: zs_formula_zs) have vars: "vars = zs_numlist vs" using vs numlist_wf_def numlist_zs_numlist zs_numlist_ex1 by blast then have wf_vs: "numlist_wf vs" using numlist_wf_def vs by auto have "verify_sat zs = (if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then [\] else [])" using * bs_second wf_xs wf_vs by simp then show ?thesis using \ vars by simp qed text \ Finally we show that for every string $x$ representing a satisfiable CNF formula there is a list of numbers representing a satisfying assignment and represented by a string of length at most $|x|$. That means there is always a string of exactly the length of $x$ consisting of a satisfying assignment plus some padding symbols. \ lemma nllength_remove1: assumes "n \ set ns" shows "nllength (n # remove1 n ns) = nllength ns" using assms nllength sum_list_map_remove1[of n ns "\n. Suc (nlength n)"] by simp lemma nllength_distinct_le: assumes "distinct ns" and "set ns \ set ms" shows "nllength ns \ nllength ms" using assms proof (induction ms arbitrary: ns) case Nil then show ?case by simp next case (Cons m ms) show ?case proof (cases "m \ set ns") case True let ?ns = "remove1 m ns" have "set ?ns \ set ms" using Cons by auto moreover have "distinct ?ns" using Cons by simp ultimately have *: "nllength ?ns \ nllength ms" using Cons by simp have "nllength ns = nllength (m # ?ns)" using True nllength_remove1 by simp also have "... = Suc (nlength m) + nllength ?ns" using nllength by simp also have "... \ Suc (nlength m) + nllength ms" using * by simp also have "... = nllength (m # ms)" using nllength by simp finally show ?thesis . next case False then have "set ns \ set ms" using Cons by auto moreover have "distinct ns" using Cons by simp ultimately have "nllength ns \ nllength ms" using Cons by simp then show ?thesis using nllength by simp qed qed lemma nlllength_nllength_concat: "nlllength nss = nllength (concat nss) + length nss" using nlllength nllength by (induction nss) simp_all fun variable :: "literal \ nat" where "variable (Neg i) = i" | "variable (Pos i) = i" lemma sum_list_eq: "ns = ms \ sum_list ns = sum_list ms" by simp lemma nllength_clause_le: "nllength (clause_n cl) \ nllength (map variable cl)" proof - have "variable x \ literal_n x" for x by (cases x) simp_all then have *: "Suc (nlength (variable x)) \ Suc (nlength (literal_n x))" for x using nlength_mono by simp let ?ns = "map literal_n cl" have "nllength (clause_n cl) = nllength ?ns" using clause_n_def by presburger also have "... = (\n\?ns. Suc (nlength n))" using nllength by simp also have "... = (\x\cl. Suc (nlength (literal_n x)))" by (smt (verit, del_insts) length_map nth_equalityI nth_map) also have "... \ (\x\cl. Suc (nlength (variable x)))" using * by (simp add: sum_list_mono) finally have "(\x\cl. Suc (nlength (variable x))) \ nllength (clause_n cl)" by simp moreover have "(\x\cl. Suc (nlength (variable x))) = nllength (map variable cl)" proof - have 1: "map (\x. Suc (nlength (variable x))) cl = map (\n. Suc (nlength n)) (map variable cl)" by simp then have "(\x\cl. Suc (nlength (variable x))) = (\n\(map variable cl). Suc (nlength n))" using sum_list_eq[OF 1] by simp then show ?thesis using nllength by simp qed ultimately show ?thesis by simp qed lemma nlllength_formula_ge: "nlllength (formula_n \) \ nlllength (map (map variable) \)" proof (induction \) case Nil then show ?case by simp next case (Cons cl \) have "nlllength (map (map variable) (cl # \)) = nlllength (map (map variable) [cl]) + nlllength (map (map variable) \)" using nlllength by simp also have "... = Suc (nllength (map variable cl)) + nlllength (map (map variable) \)" using nlllength by simp also have "... \ Suc (nllength (map variable cl)) + nlllength (formula_n \)" using Cons by simp also have "... \ Suc (nllength (clause_n cl)) + nlllength (formula_n \)" using nllength_clause_le by simp also have "... = nlllength (formula_n (cl # \))" using nlllength by (simp add: formula_n_def) finally show ?case . qed lemma variables_shorter_formula: fixes \ :: formula and vars :: "nat list" assumes "set vars \ variables \" and "distinct vars" shows "nllength vars \ nlllength (formula_n \)" proof - let ?nss = "map (map variable) \" have "nllength (concat ?nss) \ nlllength ?nss" using nlllength_nllength_concat by simp then have *: "nllength (concat ?nss) \ nlllength (formula_n \)" using nlllength_formula_ge by (meson le_trans) have "set vars \ set (concat ?nss)" proof fix n :: nat assume "n \ set vars" then have "n \ variables \" using assms(1) by auto then obtain c where c: "c \ set \" "Neg n \ set c \ Pos n \ set c" using variables_def by auto then obtain x where x: "x \ set c" "variable x = n" by auto then show "n \ set (concat (map (map variable) \))" using c by auto qed then have "nllength vars \ nllength (concat ?nss)" using nllength_distinct_le assms(2) by simp then show ?thesis using * by simp qed lemma ex_assignment_linear_length: assumes "satisfiable \" shows "\vars. (\v. v \ set vars) \ \ \ nllength vars \ nlllength (formula_n \)" proof - obtain \ where \: "\ \ \" using assms satisfiable_def by auto define poss where "poss = {v. v \ variables \ \ \ v}" then have "finite poss" using finite_variables by simp let ?beta = "\v. v \ poss" have sat: "?beta \ \" unfolding satisfies_def proof fix c :: clause assume "c \ set \" then have "satisfies_clause \ c" using satisfies_def \ by simp then obtain x where x: "x \ set c" "satisfies_literal \ x" using satisfies_clause_def by auto show "satisfies_clause ?beta c" proof (cases x) case (Neg n) then have "\ \ n" using x(2) by simp then have "n \ poss" using poss_def by simp then have "\ ?beta n" by simp then have "satisfies_literal ?beta x" using Neg by simp then show ?thesis using satisfies_clause_def x(1) by auto next case (Pos n) then have "\ n" using x(2) by simp then have "n \ poss" using poss_def Pos \c \ set \\ variables_def x(1) by auto then have "?beta n" by simp then have "satisfies_literal ?beta x" using Pos by simp then show ?thesis using satisfies_clause_def x(1) by auto qed qed obtain vars where vars: "set vars = poss" "distinct vars" using `finite poss` by (meson finite_distinct_list) moreover from this have "set vars \ variables \" using poss_def by simp ultimately have "nllength vars \ nlllength (formula_n \)" using variables_shorter_formula by simp moreover have "(\v. v \ set vars) \ \" using vars(1) sat by simp ultimately show ?thesis by auto qed lemma ex_witness_linear_length: fixes \ :: formula assumes "satisfiable \" shows "\us. bit_symbols us \ length us = length (formula_to_string \) \ verify_sat \formula_to_string \; symbols_to_string us\ = [\]" proof - obtain vars where vars: "(\v. v \ set vars) \ \" "nllength vars \ nlllength (formula_n \)" using assms ex_assignment_linear_length by auto define pad where "pad = nlllength (formula_n \) - nllength vars" then have "nllength vars + pad = nlllength (formula_n \)" using vars(2) by simp moreover define us where "us = numlist vars @ replicate pad \" ultimately have "length us = nlllength (formula_n \)" by (simp add: nllength_def) then have "length (binencode us) = length (formula_to_string \)" (is "length ?us = _") by (simp add: nlllength_def) moreover have "verify_sat \formula_to_string \; symbols_to_string ?us\ = [\]" using us_def vars(1) assms verify_sat_wf by simp moreover have "bit_symbols ?us" proof - have "binencodable (numlist vars)" using proper_symbols_numlist symbols_lt_numlist by (metis Suc_leI lessI less_Suc_numeral numeral_2_eq_2 numeral_le_iff numeral_less_iff order_less_le_trans pred_numeral_simps(3) semiring_norm(74)) moreover have "binencodable (replicate pad \)" by simp ultimately have "binencodable us" using us_def binencodable_append by simp then show ?thesis using bit_symbols_binencode by simp qed ultimately show ?thesis by blast qed lemma bit_symbols_verify_sat: "bit_symbols (verify_sat zs)" unfolding verify_sat_def Let_def by simp subsection \A Turing machine for verifying formulas\ text \ The core of the function @{const verify_sat} is the expression @{term " (\v. v \ set (zs_numlist vs)) \ zs_formula xs"}, which checks if an assignment represented by a list of variable indices satisfies a CNF formula represented by a list of lists of literals. In this section we devise a Turing machine performing this check. Recall that the numbers 0 and 1 are represented by the empty symbol sequence and the symbol sequence \textbf{1}, respectively. The Turing machines in this section are described in terms of numbers. We start with a Turing machine that checks a clause. The TM accepts on tape $j_1$ a list of numbers representing an assignment $\alpha$ and on tape $j_2$ a list of numbers representing a clause. It outputs on tape $j_3$ the number 1 if $\alpha$ satisfies the clause, and otherwise 0. To do this the TM iterates over all literals in the clause and determines the underlying variable and the sign of the literal. If the literal is positive and the variable is in the list representing $\alpha$ or if the literal is negative and the variable is not in the list, the number 1 is written to the tape $j_3$. Otherwise the tape remains unchanged. We assume $j_3$ is initialized with 0, and so it will be 1 if and only if at least one literal is satisfied by $\alpha$. The TM requires five auxiliary tapes $j_3 + 1, \dots, j_3 + 5$. Tape $j_3 + 1$ stores the literals one at a time, and later the variable; tape $j_3 + 2$ stores the sign of the literal; tape $j_3 + 3$ stores whether the variable is contained in $\alpha$; tapes $j_3 + 4$ and $j_3 + 5$ are the auxiliary tapes of @{const tm_contains}. \ definition tm_sat_clause :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_sat_clause j1 j2 j3 \ WHILE [] ; \rs. rs ! j2 \ \ DO tm_nextract 4 j2 (j3 + 1) ;; tm_mod2 (j3 + 1) (j3 + 2) ;; tm_div2 (j3 + 1) ;; tm_contains j1 (j3 + 1) (j3 + 3) ;; IF \rs. rs ! (j3 + 3) = \ \ rs ! (j3 + 2) = \ \ rs ! (j3 + 3) \ \ \ rs ! (j3 + 2) \ \ THEN tm_setn j3 1 ELSE [] ENDIF ;; tm_setn (j3 + 1) 0 ;; tm_setn (j3 + 2) 0 ;; tm_setn (j3 + 3) 0 DONE ;; tm_cr j2" lemma tm_sat_clause_tm: assumes "k \ 2" and "G \ 5" and "j3 + 5 < k" "0 < j1" "j1 < k" "j2 < k" "j1 < j3" shows "turing_machine k G (tm_sat_clause j1 j2 j3)" using tm_sat_clause_def tm_mod2_tm tm_div2_tm tm_nextract_tm tm_setn_tm tm_contains_tm Nil_tm tm_cr_tm assms turing_machine_loop_turing_machine turing_machine_branch_turing_machine by simp locale turing_machine_sat_clause = fixes j1 j2 j3 :: tapeidx begin definition "tmL1 \ tm_nextract 4 j2 (j3 + 1)" definition "tmL2 \ tmL1 ;; tm_mod2 (j3 + 1) (j3 + 2)" definition "tmL3 \ tmL2 ;; tm_div2 (j3 + 1)" definition "tmL4 \ tmL3 ;; tm_contains j1 (j3 + 1) (j3 + 3)" definition "tmI \ IF \rs. rs ! (j3 + 3) = \ \ rs ! (j3 + 2) = \ \ rs ! (j3 + 3) \ \ \ rs ! (j3 + 2) \ \ THEN tm_setn j3 1 ELSE [] ENDIF" definition "tmL5 \ tmL4 ;; tmI" definition "tmL6 \ tmL5 ;; tm_setn (j3 + 1) 0" definition "tmL7 \ tmL6 ;; tm_setn (j3 + 2) 0" definition "tmL8 \ tmL7 ;; tm_setn (j3 + 3) 0" definition "tmL \ WHILE [] ; \rs. rs ! j2 \ \ DO tmL8 DONE" definition "tm2 \ tmL ;; tm_cr j2" lemma tm2_eq_tm_sat_clause: "tm2 = tm_sat_clause j1 j2 j3" unfolding tm2_def tmL_def tmL8_def tmL7_def tmL6_def tmL5_def tmL4_def tmL3_def tmI_def tmL2_def tmL1_def tm_sat_clause_def by simp context fixes tps0 :: "tape list" and k :: nat and vars :: "nat list" and clause :: clause assumes jk: "0 < j1" "j1 \ j2" "j3 + 5 < k" "j1 < j3" "j2 < j3" "0 < j2" "length tps0 = k" assumes tps0: "tps0 ! j1 = nltape' vars 0" "tps0 ! j2 = nltape' (clause_n clause) 0" "tps0 ! j3 = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 1) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 3) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 4) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 5) = (\0\\<^sub>N, 1)" begin abbreviation "sat_take t \ satisfies_clause (\v. v \ set vars) (take t clause)" definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j2 := nltape' (clause_n clause) t, j3 := (\sat_take t\\<^sub>B, 1)]" lemma tpsL0: "tpsL 0 = tps0" proof - have "nltape' (clause_n clause) 0 = tps0 ! j2" using tps0(2) by presburger moreover have "\sat_take 0\\<^sub>B = \0\\<^sub>N" using satisfies_clause_def by simp ultimately show ?thesis using tpsL_def tps0 jk by (metis list_update_id) qed definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take t\\<^sub>B, 1), j3 + 1 := (\literal_n (clause ! t)\\<^sub>N, 1)]" lemma tmL1 [transforms_intros]: assumes "ttt = 12 + 2 * nlength (clause_n clause ! t)" and "t < length (clause_n clause)" shows "transforms tmL1 (tpsL t) ttt (tpsL1 t)" unfolding tmL1_def proof (tform tps: assms tps0 tpsL_def tpsL1_def jk) have len: "t < length clause" using assms(2) clause_n_def by simp show "ttt = 12 + 2 * nlength 0 + 2 * nlength (clause_n clause ! t)" using assms(1) by simp have *: "j2 \ j3" using jk by simp have **: "clause_n clause ! t = literal_n (clause ! t)" using len by (simp add: clause_n_def) show "tpsL1 t = (tpsL t) [j2 := nltape' (clause_n clause) (Suc t), j3 + 1 := (\clause_n clause ! t\\<^sub>N, 1)]" unfolding tpsL_def tpsL1_def using list_update_swap[OF *, of tps0] by (simp add: **) qed definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take t\\<^sub>B, 1), j3 + 1 := (\literal_n (clause ! t)\\<^sub>N, 1), j3 + 2 := (\literal_n (clause ! t) mod 2\\<^sub>N, 1)]" lemma tmL2 [transforms_intros]: assumes "ttt = 12 + 2 * nlength (clause_n clause ! t) + 1" and "t < length (clause_n clause)" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" unfolding tmL2_def by (tform tps: assms tps0 tpsL2_def tpsL1_def jk) definition tpsL3 :: "nat \ tape list" where "tpsL3 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take t\\<^sub>B, 1), j3 + 1 := (\literal_n (clause ! t) div 2\\<^sub>N, 1), j3 + 2 := (\literal_n (clause ! t) mod 2\\<^sub>N, 1)]" lemma tmL3 [transforms_intros]: assumes "ttt = 16 + 4 * nlength (clause_n clause ! t)" and "t < length (clause_n clause)" shows "transforms tmL3 (tpsL t) ttt (tpsL3 t)" unfolding tmL3_def proof (tform tps: assms(2) tps0 tpsL3_def tpsL2_def jk) have len: "t < length clause" using assms(2) clause_n_def by simp have **: "clause_n clause ! t = literal_n (clause ! t)" using len by (simp add: clause_n_def) show "ttt = 12 + 2 * nlength (clause_n clause ! t) + 1 + (2 * nlength (literal_n (clause ! t)) + 3)" using assms(1) ** by simp qed definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take t\\<^sub>B, 1), j3 + 1 := (\literal_n (clause ! t) div 2\\<^sub>N, 1), j3 + 2 := (\literal_n (clause ! t) mod 2\\<^sub>N, 1), j3 + 3 := (\literal_n (clause ! t) div 2 \ set vars\\<^sub>B, 1)]" lemma tmL4 [transforms_intros]: assumes "ttt = 20 + 4 * nlength (clause_n clause ! t) + 67 * (nllength vars)\<^sup>2" and "t < length (clause_n clause)" shows "transforms tmL4 (tpsL t) ttt (tpsL4 t)" unfolding tmL4_def proof (tform tps: assms(2) tps0 tpsL4_def tpsL3_def jk time: assms(1)) have "tpsL3 t ! (j3 + 4) = (\0\\<^sub>N, 1)" using tpsL3_def tps0 jk by simp then show "tpsL3 t ! (j3 + 3 + 1) = (\0\\<^sub>N, 1)" by (metis ab_semigroup_add_class.add_ac(1) numeral_plus_one semiring_norm(2) semiring_norm(8)) have "tpsL3 t ! (j3 + 5) = (\0\\<^sub>N, 1)" using tpsL3_def tps0 jk by simp then show "tpsL3 t ! (j3 + 3 + 2) = (\0\\<^sub>N, 1)" by (simp add: numeral_Bit1) qed definition tpsL5 :: "nat \ tape list" where "tpsL5 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\literal_n (clause ! t) div 2\\<^sub>N, 1), j3 + 2 := (\literal_n (clause ! t) mod 2\\<^sub>N, 1), j3 + 3 := (\literal_n (clause ! t) div 2 \ set vars\\<^sub>B, 1)]" lemma tmI [transforms_intros]: assumes "ttt = 16" and "t < length (clause_n clause)" shows "transforms tmI (tpsL4 t) ttt (tpsL5 t)" unfolding tmI_def proof (tform tps: jk tpsL4_def time: assms(1)) show "10 + 2 * nlength (if sat_take t then 1 else 0) + 2 * nlength 1 + 2 \ ttt" using assms(1) nlength_0 nlength_1_simp by simp have len: "t < length clause" using assms(2) by (simp add: clause_n_def) let ?l = "clause ! t" have 1: "read (tpsL4 t) ! (j3 + 3) = \ \ literal_n ?l div 2 \ set vars" using tpsL4_def jk read_ncontents_eq_0[of "tpsL4 t" "j3 + 3"] by simp have 2: "read (tpsL4 t) ! (j3 + 2) = \ \ literal_n ?l mod 2 = 0" using tpsL4_def jk read_ncontents_eq_0[of "tpsL4 t" "j3 + 2"] by simp let ?a = "\v. v \ set vars" let ?cond = "read (tpsL4 t) ! (j3 + 3) = \ \ read (tpsL4 t) ! (j3 + 2) = \ \ read (tpsL4 t) ! (j3 + 3) \ \ \ read (tpsL4 t) ! (j3 + 2) \ \" have *: "?cond \ satisfies_literal ?a ?l" proof (cases ?l) case (Neg v) then have "literal_n ?l div 2 = v" "literal_n ?l mod 2 = 0" by simp_all moreover from this have "satisfies_literal ?a ?l \ v \ set vars" using Neg by simp ultimately show ?thesis using 1 2 by simp next case (Pos v) then have "literal_n ?l div 2 = v" "literal_n ?l mod 2 = 1" by simp_all moreover from this have "satisfies_literal ?a ?l \ v \ set vars" using Pos by simp ultimately show ?thesis using 1 2 by simp qed have **: "sat_take (Suc t) \ sat_take t \ satisfies_literal ?a ?l" using satisfies_clause_take[OF len] by simp show "tpsL5 t = (tpsL4 t)[j3 := (\1\\<^sub>N, 1)]" if ?cond proof - have "(if sat_take (Suc t) then 1::nat else 0) = 1" using that * ** by simp then show ?thesis unfolding tpsL5_def tpsL4_def using that by (simp add: list_update_swap) qed show "tpsL5 t = (tpsL4 t)" if "\ ?cond" proof - have "sat_take t = sat_take (Suc t)" using * ** that by simp then show ?thesis unfolding tpsL5_def tpsL4_def using that by (simp add: list_update_swap) qed qed lemma tmL5 [transforms_intros]: assumes "ttt = 36 + 4 * nlength (clause_n clause ! t) + 67 * (nllength vars)\<^sup>2" and "t < length (clause_n clause)" shows "transforms tmL5 (tpsL t) ttt (tpsL5 t)" unfolding tmL5_def by (tform tps: assms) definition tpsL6 :: "nat \ tape list" where "tpsL6 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\0\\<^sub>N, 1), j3 + 2 := (\literal_n (clause ! t) mod 2\\<^sub>N, 1), j3 + 3 := (\literal_n (clause ! t) div 2 \ set vars\\<^sub>B, 1)]" lemma tmL6 [transforms_intros]: assumes "ttt = 46 + 4 * nlength (clause_n clause ! t) + 67 * (nllength vars)\<^sup>2 + 2 * nlength (literal_n (clause ! t) div 2)" and "t < length (clause_n clause)" shows "transforms tmL6 (tpsL t) ttt (tpsL6 t)" unfolding tmL6_def by (tform tps: assms tps0 tpsL6_def tpsL5_def jk) definition tpsL7 :: "nat \ tape list" where "tpsL7 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\0\\<^sub>N, 1), j3 + 2 := (\0\\<^sub>N, 1), j3 + 3 := (\literal_n (clause ! t) div 2 \ set vars\\<^sub>B, 1)]" lemma tmL7 [transforms_intros]: assumes "ttt = 56 + 4 * nlength (clause_n clause ! t) + 67 * (nllength vars)\<^sup>2 + 2 * nlength (literal_n (clause ! t) div 2) + 2 * nlength (literal_n (clause ! t) mod 2)" and "t < length (clause_n clause)" shows "transforms tmL7 (tpsL t) ttt (tpsL7 t)" unfolding tmL7_def by (tform tps: assms tps0 tpsL7_def tpsL6_def jk) definition tpsL8 :: "nat \ tape list" where "tpsL8 t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\0\\<^sub>N, 1), j3 + 2 := (\0\\<^sub>N, 1), j3 + 3 := (\0\\<^sub>N, 1)]" lemma tmL8: assumes "ttt = 66 + 4 * nlength (clause_n clause ! t) + 67 * (nllength vars)\<^sup>2 + 2 * nlength (literal_n (clause ! t) div 2) + 2 * nlength (literal_n (clause ! t) mod 2) + 2 * nlength (if literal_n (clause ! t) div 2 \ set vars then 1 else 0)" and "t < length (clause_n clause)" shows "transforms tmL8 (tpsL t) ttt (tpsL8 t)" unfolding tmL8_def by (tform tps: assms tps0 tpsL8_def tpsL7_def jk) lemma tmL8': assumes "ttt = 70 + 6 * nllength (clause_n clause) + 67 * (nllength vars)\<^sup>2" and "t < length (clause_n clause)" shows "transforms tmL8 (tpsL t) ttt (tpsL8 t)" proof - let ?l = "literal_n (clause ! t)" let ?ll = "clause_n clause ! t" let ?t = "66 + 4 * nlength ?ll + 67 * (nllength vars)\<^sup>2 + 2 * nlength (?l div 2) + 2 * nlength (?l mod 2) + 2 * nlength (if ?l div 2 \ set vars then 1 else 0)" have "?t = 66 + 4 * nlength ?ll + 67 * (nllength vars)\<^sup>2 + 2 * nlength (?ll div 2) + 2 * nlength (?ll mod 2) + 2 * nlength (if ?ll div 2 \ set vars then 1 else 0)" using assms(2) clause_n_def by simp also have "... \ 66 + 4 * nlength ?ll + 67 * (nllength vars)\<^sup>2 + 2 * nlength ?ll + 2 * nlength (?ll mod 2) + 2 * nlength (if ?ll div 2 \ set vars then 1 else 0)" using nlength_mono[of "?ll div 2" "?ll"] by simp also have "... = 66 + 6 * nlength ?ll + 67 * (nllength vars)\<^sup>2 + 2 * nlength (?ll mod 2) + 2 * nlength (if ?ll div 2 \ set vars then 1 else 0)" by simp also have "... \ 66 + 6 * nlength ?ll + 67 * (nllength vars)\<^sup>2 + 2 * nlength 1 + 2 * nlength (if ?ll div 2 \ set vars then 1 else 0)" using nlength_mono by simp also have "... \ 66 + 6 * nlength ?ll + 67 * (nllength vars)\<^sup>2 + 2 * nlength 1 + 2 * nlength 1" using nlength_mono by simp also have "... = 70 + 6 * nlength ?ll + 67 * (nllength vars)\<^sup>2" using nlength_1_simp by simp also have "... \ 70 + 6 * nllength (clause_n clause) + 67 * (nllength vars)\<^sup>2" using assms(2) member_le_nllength by simp finally have "?t \ ttt" using assms(1) by simp then show ?thesis using assms tmL8 transforms_monotone by blast qed definition tpsL8' :: "nat \ tape list" where "tpsL8' t \ tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1)]" lemma tpsL8': "tpsL8' = tpsL8" proof - { fix t :: nat have "tpsL8 t = tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\0\\<^sub>N, 1), j3 + 2 := (\0\\<^sub>N, 1)]" unfolding tpsL8_def using tps0 list_update_id[of "tps0" "j3 + 3"] jk by (simp add: list_update_swap[of _ "j3 + 3"]) also have "... = tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\0\\<^sub>N, 1)]" unfolding tpsL8_def using tps0 list_update_id[of "tps0" "j3 + 2"] jk by (simp add: list_update_swap[of _ "Suc (Suc j3)"]) also have "... = tps0 [j2 := nltape' (clause_n clause) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1)]" unfolding tpsL8_def using tps0 list_update_id[of "tps0" "j3 + 1"] jk by (simp add: list_update_swap[of _ "Suc j3"]) also have "... = tpsL8' t" using tpsL8'_def by simp finally have "tpsL8 t = tpsL8' t" . } then show ?thesis by auto qed lemma tmL8'' [transforms_intros]: assumes "ttt = 70 + 6 * nllength (clause_n clause) + 67 * (nllength vars)\<^sup>2" and "t < length (clause_n clause)" shows "transforms tmL8 (tpsL t) ttt (tpsL8' t)" using tmL8' tpsL8' assms by simp lemma tmL [transforms_intros]: assumes "ttt = length (clause_n clause) * (72 + 6 * nllength (clause_n clause) + 67 * (nllength vars)\<^sup>2) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length (clause_n clause)))" unfolding tmL_def proof (tform) let ?t = "70 + 6 * nllength (clause_n clause) + 67 * (nllength vars)\<^sup>2" have "tpsL8' t = tpsL (Suc t)" for t using tpsL8'_def tpsL_def by simp then show "\i. i < length (clause_n clause) \ transforms tmL8 (tpsL i) ?t (tpsL (Suc i))" using tmL8'' by simp let ?ns = "clause_n clause" have *: "tpsL t ! j2 = nltape' ?ns t" for t using tpsL_def jk by simp moreover have "read (tpsL t) ! j2 = tpsL t :.: j2" for t using tapes_at_read'[of j2 "tpsL t"] tpsL_def jk by simp ultimately have "read (tpsL t) ! j2 = |.| (nltape' ?ns t)" for t by simp then have "read (tpsL t) ! j2 = \ \ (t \ length ?ns)" for t using nltape'_tape_read by simp then show "\i. i < length ?ns \ read (tpsL i) ! j2 \ \" "\ read (tpsL (length ?ns)) ! j2 \ \" using * by simp_all show "length ?ns * (70 + 6 * nllength ?ns + 67 * (nllength vars)\<^sup>2 + 2) + 1 \ ttt" using assms by simp qed definition tps1 :: "tape list" where "tps1 \ tps0 [j2 := nltape' (clause_n clause) (length (clause_n clause)), j3 := (\satisfies_clause (\v. v \ set vars) clause\\<^sub>B, 1)]" lemma tps1: "tps1 = tpsL (length (clause_n clause))" proof - have "length (clause_n clause) = length clause" by (simp add: clause_n_def) then show ?thesis using tps1_def tpsL_def by simp qed lemma tm1 [transforms_intros]: assumes "ttt = length (clause_n clause) * (72 + 6 * nllength (clause_n clause) + 67 * (nllength vars)\<^sup>2) + 1" shows "transforms tmL tps0 ttt tps1" using tmL tpsL0 assms tps1 by simp definition tps2 :: "tape list" where "tps2 \ tps0 [j2 := nltape' (clause_n clause) 0, j3 := (\satisfies_clause (\v. v \ set vars) clause\\<^sub>B, 1)]" lemma tm2: assumes "ttt = length (clause_n clause) * (72 + 6 * nllength (clause_n clause) + 67 * (nllength vars)\<^sup>2) + nllength (clause_n clause) + 4" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: assms tps0 tps1_def jk) have *: "tps1 ! j2 = nltape' (clause_n clause) (length (clause_n clause))" using tps1_def jk by simp then show "clean_tape (tps1 ! j2)" using clean_tape_nlcontents by simp have neq: "j3 \ j2" using jk by simp have "tps2 = tps1[j2 := nltape' (clause_n clause) 0]" unfolding tps2_def tps1_def by (simp add: list_update_swap[OF neq]) moreover have "tps1 ! j2 |#=| 1 = nltape' (clause_n clause) 0" using * by simp ultimately show "tps2 = tps1[j2 := tps1 ! j2 |#=| 1]" by simp qed definition tps2' :: "tape list" where "tps2' \ tps0 [j3 := (\satisfies_clause (\v. v \ set vars) clause\\<^sub>B, 1)]" lemma tm2': assumes "ttt = 79 * (nllength (clause_n clause)) ^ 2 + 67 * (nllength (clause_n clause)) * nllength vars ^ 2 + 4" shows "transforms tm2 tps0 ttt tps2'" proof - let ?l = "nllength (clause_n clause)" let ?t = "length (clause_n clause) * (72 + 6 * ?l + 67 * (nllength vars)\<^sup>2) + ?l + 4" have "?t \ ?l * (72 + 6 * ?l + 67 * (nllength vars)\<^sup>2) + ?l + 4" by (simp add: length_le_nllength) also have "... = ?l * (73 + 6 * ?l + 67 * (nllength vars)\<^sup>2) + 4" by algebra also have "... = 73 * ?l + 6 * ?l ^ 2 + 67 * ?l * (nllength vars)\<^sup>2 + 4" by algebra also have "... \ 79 * ?l ^ 2 + 67 * ?l * (nllength vars)\<^sup>2 + 4" using linear_le_pow by simp finally have "?t \ ttt" using assms by simp moreover have "tps2' = tps2" unfolding tps2'_def tps2_def using jk tps0 by (metis tape_list_eq) ultimately show ?thesis using tps2'_def tm2 assms transforms_monotone by simp qed end (* context *) end (* locale *) lemma transforms_tm_sat_clauseI [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and ttt k :: nat and vars :: "nat list" and clause :: "literal list" assumes "0 < j1" "j1 \ j2" "j3 + 5 < k" "j1 < j3" "j2 < j3" "0 < j2" "length tps = k" assumes "tps ! j1 = nltape' vars 0" "tps ! j2 = nltape' (clause_n clause) 0" "tps ! j3 = (\0\\<^sub>N, 1)" "tps ! (j3 + 1) = (\0\\<^sub>N, 1)" "tps ! (j3 + 2) = (\0\\<^sub>N, 1)" "tps ! (j3 + 3) = (\0\\<^sub>N, 1)" "tps ! (j3 + 4) = (\0\\<^sub>N, 1)" "tps ! (j3 + 5) = (\0\\<^sub>N, 1)" assumes "tps' = tps [j3 := (\satisfies_clause (\v. v \ set vars) clause\\<^sub>B, 1)]" assumes "ttt = 79 * (nllength (clause_n clause)) ^ 2 + 67 * (nllength (clause_n clause)) * nllength vars ^ 2 + 4" shows "transforms (tm_sat_clause j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_sat_clause j1 j2 j3 . show ?thesis using assms loc.tps2'_def loc.tm2' loc.tm2_eq_tm_sat_clause by simp qed text \ The following Turing machine expects a list of lists of numbers representing a formula $\phi$ on tape $j_1$ and a list of numbers representing an assignment $\alpha$ on tape $j_2$. It outputs on tape $j_3$ the number 1 if $\alpha$ satisfies $\phi$, and otherwise the number 0. To do so the TM iterates over all clauses in $\phi$ and uses @{const tm_sat_clause} on each of them. It requires seven auxiliary tapes: $j_3 + 1$ to store the clauses one at a time, $j_3 + 2$ to store the results of @{const tm_sat_clause}, whose auxiliary tapes are $j_3 + 3, \dots, j_3 + 7$. \ definition tm_sat_formula :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_sat_formula j1 j2 j3 \ tm_setn j3 1 ;; WHILE [] ; \rs. rs ! j1 \ \ DO tm_nextract \ j1 (j3 + 1) ;; tm_sat_clause j2 (j3 + 1) (j3 + 2) ;; IF \rs. rs ! (j3 + 2) = \ THEN tm_setn j3 0 ELSE [] ENDIF ;; tm_erase_cr (j3 + 1) ;; tm_setn (j3 + 2) 0 DONE" lemma tm_sat_formula_tm: assumes "k \ 2" and "G \ 6" and "0 < j1" "j1 \ j2" "j3 + 7 < k" "j1 < j3" "j2 < j3" "0 < j2" shows "turing_machine k G (tm_sat_formula j1 j2 j3)" using tm_sat_formula_def tm_sat_clause_tm tm_nextract_tm tm_setn_tm assms Nil_tm tm_erase_cr_tm turing_machine_loop_turing_machine turing_machine_branch_turing_machine by simp locale turing_machine_sat_formula = fixes j1 j2 j3 :: tapeidx begin definition "tm1 \ tm_setn j3 1" definition "tmL1 \ tm_nextract \ j1 (j3 + 1)" definition "tmL2 \ tmL1 ;; tm_sat_clause j2 (j3 + 1) (j3 + 2)" definition "tmI \ IF \rs. rs ! (j3 + 2) = \ THEN tm_setn j3 0 ELSE [] ENDIF" definition "tmL3 \ tmL2 ;; tmI" definition "tmL4 \ tmL3 ;; tm_erase_cr (j3 + 1)" definition "tmL5 \ tmL4 ;; tm_setn (j3 + 2) 0" definition "tmL \ WHILE [] ; \rs. rs ! j1 \ \ DO tmL5 DONE" definition "tm2 \ tm1 ;; tmL" lemma tm2_eq_tm_sat_formula: "tm2 = tm_sat_formula j1 j2 j3" unfolding tm2_def tm1_def tmL_def tmL5_def tmL4_def tmL3_def tmI_def tmL2_def tmL1_def tm_sat_formula_def by simp context fixes tps0 :: "tape list" and k :: nat and vars :: "nat list" and \ :: formula assumes jk: "0 < j1" "j1 \ j2" "j3 + 7 < k" "j1 < j3" "j2 < j3" "0 < j2" "length tps0 = k" assumes tps0: "tps0 ! j1 = nlltape' (formula_n \) 0" "tps0 ! j2 = nltape' vars 0" "tps0 ! j3 = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 1) = (\[]\\<^sub>N\<^sub>L, 1)" "tps0 ! (j3 + 2) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 3) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 4) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 5) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 6) = (\0\\<^sub>N, 1)" "tps0 ! (j3 + 7) = (\0\\<^sub>N, 1)" begin definition "tps1 \ tps0 [j3 := (\1\\<^sub>N, 1)]" lemma tm1 [transforms_intros]: assumes "ttt = 12" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps0 tps1_def jk) show "ttt = 10 + 2 * nlength 0 + 2 * nlength 1" using assms nlength_1_simp by simp qed abbreviation "sat_take t \ (\v. v \ set vars) \ take t \" definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j1 := nlltape' (formula_n \) t, j3 := (\sat_take t\\<^sub>B, 1)]" lemma tpsL0: "tpsL 0 = tps1" proof - have "nlltape' (formula_n \) 0 = tps1 ! j1" using tps0(1) tps1_def jk by simp moreover have "\sat_take 0\\<^sub>B = \1\\<^sub>N" using satisfies_def by simp ultimately show ?thesis using tpsL_def tps0 jk tps1_def by (metis list_update_id) qed definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take t\\<^sub>B, 1), j3 + 1 := (\formula_n \ ! t\\<^sub>N\<^sub>L, 1)]" lemma tmL1 [transforms_intros]: assumes "ttt = 12 + 2 * nllength (formula_n \ ! t)" and "t < length (formula_n \)" shows "transforms tmL1 (tpsL t) ttt (tpsL1 t)" unfolding tmL1_def proof (tform tps: assms tps0 tpsL_def tpsL1_def jk) show "ttt = 12 + 2 * nllength [] + 2 * nllength (formula_n \ ! t)" using assms(1) by simp show "tpsL1 t = (tpsL t) [j1 := nlltape' (formula_n \) (Suc t), j3 + 1 := (\formula_n \ ! t\\<^sub>N\<^sub>L, 1)]" using tpsL1_def tpsL_def jk by (simp add: list_update_swap) qed definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take t\\<^sub>B, 1), j3 + 1 := (\formula_n \ ! t\\<^sub>N\<^sub>L, 1), j3 + 2 := (\satisfies_clause (\v. v \ set vars) (\ ! t)\\<^sub>B, 1)]" lemma tmL2 [transforms_intros]: assumes "ttt = 12 + 2 * nllength (formula_n \ ! t) + (79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2 + 4)" and "t < length (formula_n \)" shows "transforms tmL2 (tpsL t) ttt (tpsL2 t)" unfolding tmL2_def proof (tform tps: assms tps0 tpsL_def tpsL1_def jk) let ?clause = "\ ! t" have *: "formula_n \ ! t = clause_n ?clause" using assms(2) formula_n_def by simp then have "(\formula_n \ ! t\\<^sub>N\<^sub>L, 1) = nltape' (clause_n ?clause) 0" by simp then show "tpsL1 t ! (j3 + 1) = nltape' (clause_n ?clause) 0" using tpsL1_def jk by simp have "j3 + 2 + 1 = j3 + 3" by simp moreover have "tpsL1 t ! (j3 + 3) = (\0\\<^sub>N, 1)" using tpsL1_def tps0 jk by simp ultimately show "tpsL1 t ! (j3 + 2 + 1) = (\0\\<^sub>N, 1)" by metis have "j3 + 2 + 2 = j3 + 4" by simp moreover have "tpsL1 t ! (j3 + 4) = (\0\\<^sub>N, 1)" using tpsL1_def tps0 jk by simp ultimately show "tpsL1 t ! (j3 + 2 + 2) = (\0\\<^sub>N, 1)" by metis have "j3 + 2 + 3 = j3 + 5" by simp moreover have "tpsL1 t ! (j3 + 5) = (\0\\<^sub>N, 1)" using tpsL1_def tps0 jk by simp ultimately show "tpsL1 t ! (j3 + 2 + 3) = (\0\\<^sub>N, 1)" by metis have "j3 + 2 + 4 = j3 + 6" by simp moreover have "tpsL1 t ! (j3 + 6) = (\0\\<^sub>N, 1)" using tpsL1_def tps0 jk by simp ultimately show "tpsL1 t ! (j3 + 2 + 4) = (\0\\<^sub>N, 1)" by metis have "j3 + 2 + 5 = j3 + 7" by simp moreover have "tpsL1 t ! (j3 + 7) = (\0\\<^sub>N, 1)" using tpsL1_def tps0 jk by simp ultimately show "tpsL1 t ! (j3 + 2 + 5) = (\0\\<^sub>N, 1)" by metis show "tpsL2 t = (tpsL1 t) [j3 + 2 := (\satisfies_clause (\v. v \ set vars) (\ ! t)\\<^sub>B, 1)]" unfolding tpsL2_def tpsL1_def by simp show "ttt = 12 + 2 * nllength (formula_n \ ! t) + (79 * (nllength (clause_n (\ ! t)))\<^sup>2 + 67 * nllength (clause_n (\ ! t)) * (nllength vars)\<^sup>2 + 4)" using assms(1) * by simp qed definition tpsL3 :: "nat \ tape list" where "tpsL3 t \ tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\formula_n \ ! t\\<^sub>N\<^sub>L, 1), j3 + 2 := (\satisfies_clause (\v. v \ set vars) (\ ! t)\\<^sub>B, 1)]" lemma tmI [transforms_intros]: assumes "ttt = 16" and "t < length (formula_n \)" shows "transforms tmI (tpsL2 t) ttt (tpsL3 t)" unfolding tmI_def proof (tform tps: assms(2) tps0 tpsL2_def tpsL3_def jk time: assms(1)) show "10 + 2 * nlength (if sat_take t then 1 else 0) + 2 * nlength 0 + 2 \ ttt" using assms(1) nlength_1_simp by simp let ?a = "\v. v \ set vars" let ?cl = "\ ! t" have *: "read (tpsL2 t) ! (j3 + 2) \ \ \ satisfies_clause ?a ?cl" using tpsL2_def jk read_ncontents_eq_0[of "tpsL2 t" "j3 + 2"] by force have len: "t < length \" using assms(2) by (simp add: formula_n_def) have **: "sat_take (Suc t) \ sat_take t \ satisfies_clause ?a ?cl" using satisfies_take[OF len] by simp show "tpsL3 t = (tpsL2 t)[j3 := (\0\\<^sub>N, 1)]" if "read (tpsL2 t) ! (j3 + 2) = \" proof - have "(if sat_take (Suc t) then 1::nat else 0) = 0" using that * ** by simp then show ?thesis unfolding tpsL3_def tpsL2_def using that by (simp add: list_update_swap) qed show "tpsL3 t = (tpsL2 t)" if "read (tpsL2 t) ! (j3 + 2) \ \" proof - have "sat_take t = sat_take (Suc t)" using * ** that by simp then show ?thesis unfolding tpsL3_def tpsL2_def using that by (simp add: list_update_swap) qed qed lemma tmL3 [transforms_intros]: assumes "ttt = 32 + 2 * nllength (formula_n \ ! t) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2" and "t < length (formula_n \)" shows "transforms tmL3 (tpsL t) ttt (tpsL3 t)" unfolding tmL3_def by (tform tps: assms) definition tpsL4 :: "nat \ tape list" where "tpsL4 t \ tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\[]\\<^sub>N\<^sub>L, 1), j3 + 2 := (\satisfies_clause (\v. v \ set vars) (\ ! t)\\<^sub>B, 1)]" lemma tmL4 [transforms_intros]: assumes "ttt = 39 + 4 * nllength (formula_n \ ! t) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2" and "t < length (formula_n \)" shows "transforms tmL4 (tpsL t) ttt (tpsL4 t)" unfolding tmL4_def proof (tform tps: assms(2) tps0 tpsL3_def tpsL4_def jk) let ?zs = "numlist (formula_n \ ! t)" have *: "tpsL3 t ! (j3 + 1) = (\formula_n \ ! t\\<^sub>N\<^sub>L, 1)" using tpsL3_def jk by simp then show "tpsL3 t ::: (j3 + 1) = \?zs\" using nlcontents_def by simp show "proper_symbols ?zs" using proper_symbols_numlist by simp show "tpsL4 t = (tpsL3 t)[j3 + 1 := (\[]\, 1)]" unfolding tpsL4_def tpsL3_def using nlcontents_Nil by (simp add: list_update_swap) show "ttt = 32 + 2 * nllength (formula_n \ ! t) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2 + (tpsL3 t :#: (j3 + 1) + 2 * length (numlist (formula_n \ ! t)) + 6)" using * assms(1) nllength_def by simp qed definition tpsL5 :: "nat \ tape list" where "tpsL5 t \ tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\[]\\<^sub>N\<^sub>L, 1), j3 + 2 := (\0\\<^sub>N, 1)]" lemma tmL5: assumes "ttt = 49 + 4 * nllength (formula_n \ ! t) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2 + 2 * nlength (if satisfies_clause (\v. v \ set vars) (\ ! t) then 1 else 0)" and "t < length (formula_n \)" shows "transforms tmL5 (tpsL t) ttt (tpsL5 t)" unfolding tmL5_def by (tform tps: assms tps0 tpsL4_def tpsL5_def jk) definition tpsL5' :: "nat \ tape list" where "tpsL5' t \ tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1)]" lemma tpsL5': "tpsL5' = tpsL5" proof fix t have 5: "j1 \ j3 + 1" using jk by simp have 4: "j3 \ j3 + 1" by simp have 1: "j3 \ j3 + 2" by simp have 2: "j3 + 1 \ j3 + 2" by simp have 22: "Suc j3 \ Suc (Suc j3)" by simp have 3: "j1 \ j3 + 2" using jk by simp let ?tps1 = "tps0 [j1 := nlltape' (formula_n \) (Suc t)]" let ?tps2 = "tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1)]" have "tpsL5 t = tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1), j3 + 1 := (\[]\\<^sub>N\<^sub>L, 1)]" unfolding tpsL5_def using tps0(5) list_update_swap[OF 2, of ?tps2] list_update_swap[OF 1, of ?tps1] list_update_swap[OF 3, of tps0] list_update_id[of tps0 "j3 + 2"] by (simp only:) also have "... = tps0 [j1 := nlltape' (formula_n \) (Suc t), j3 := (\sat_take (Suc t)\\<^sub>B, 1)]" using tps0(4) list_update_swap[OF 4, of ?tps1] list_update_swap[OF 5, of tps0] list_update_id[of tps0 "j3 + 1"] by (simp only:) finally show "tpsL5' t = tpsL5 t" using tpsL5'_def by simp qed lemma tmL5' [transforms_intros]: assumes "ttt = 51 + 83 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2" and "t < length (formula_n \)" shows "transforms tmL5 (tpsL t) ttt (tpsL5' t)" proof - let ?ttt = "49 + 4 * nllength (formula_n \ ! t) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2 + 2 * nlength (if satisfies_clause (\v. v \ set vars) (\ ! t) then 1 else 0)" have "?ttt \ 49 + 4 * nllength (formula_n \ ! t) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2 + 2 * nlength 1" by simp also have "... = 51 + 4 * nllength (formula_n \ ! t) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2" using nlength_1_simp by simp also have "... \ 51 + 4 * nlllength (formula_n \) + 79 * (nllength (formula_n \ ! t))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2" using member_le_nlllength_1[of "formula_n \ ! t" "formula_n \"] assms(2) by simp also have "... \ 51 + 4 * nlllength (formula_n \) + 79 * (nlllength (formula_n \))\<^sup>2 + 67 * nllength (formula_n \ ! t) * (nllength vars)\<^sup>2" using member_le_nlllength_1[of "formula_n \ ! t" "formula_n \"] assms(2) by simp also have "... \ 51 + 4 * nlllength (formula_n \) + 79 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2" using member_le_nlllength_1[of "formula_n \ ! t" "formula_n \"] assms(2) by auto also have "... \ 51 + 83 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2" using linear_le_pow by simp finally have "?ttt \ ttt" using assms(1) by simp then show ?thesis using tpsL5' transforms_monotone[OF tmL5] assms by simp qed lemma tmL [transforms_intros]: assumes "ttt = length (formula_n \) * (53 + 83 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length (formula_n \)))" unfolding tmL_def proof (tform) let ?t = "51 + 83 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2" have "tpsL5' t = tpsL (Suc t)" for t using tpsL5'_def tpsL_def by simp then show "\t. t < length (formula_n \) \ transforms tmL5 (tpsL t) ?t (tpsL (Suc t))" using tmL5' by simp let ?nss = "formula_n \" have *: "tpsL t ! j1 = nlltape' ?nss t" for t using tpsL_def jk by simp moreover have "read (tpsL t) ! j1 = tpsL t :.: j1" for t using tapes_at_read'[of j1 "tpsL t"] tpsL_def jk by simp ultimately have "read (tpsL t) ! j1 = |.| (nlltape' ?nss t)" for t by simp then have "read (tpsL t) ! j1 = \ \ (t \ length ?nss)" for t using nlltape'_tape_read by simp then show "\i. i < length ?nss \ read (tpsL i) ! j1 \ \" "\ read (tpsL (length ?nss)) ! j1 \ \" using * by simp_all show "length (formula_n \) * (?t + 2) + 1 \ ttt" using assms by simp qed lemma tm2: assumes "ttt = length (formula_n \) * (53 + 83 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2) + 13" shows "transforms tm2 tps0 ttt (tpsL (length (formula_n \)))" unfolding tm2_def proof (tform tps: assms tps0 tpsL4_def tpsL5_def jk tpsL0) show "transforms tm1 tps0 12 (tpsL 0)" using tm1 tpsL0 by simp qed definition tps2 :: "tape list" where "tps2 \ tps0 [j1 := nlltape (formula_n \), j3 := (\(\v. v \ set vars) \ \\\<^sub>B, 1)]" lemma tps2: "tps2 = tpsL (length (formula_n \))" using formula_n_def tps2_def tpsL_def by simp lemma tm2': assumes "ttt = length (formula_n \) * (53 + 83 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2) + 13" shows "transforms tm2 tps0 ttt tps2" using tm2 tps2 assms by simp end (* context *) end (* locale *) lemma transforms_tm_sat_formulaI [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and ttt k :: nat and vars :: "nat list" and \ :: formula assumes "0 < j1" "j1 \ j2" "j3 + 7 < k" "j1 < j3" "j2 < j3" "0 < j2" "length tps = k" assumes "tps ! j1 = nlltape' (formula_n \) 0" "tps ! j2 = nltape' vars 0" "tps ! j3 = (\0\\<^sub>N, 1)" "tps ! (j3 + 1) = (\[]\\<^sub>N\<^sub>L, 1)" "tps ! (j3 + 2) = (\0\\<^sub>N, 1)" "tps ! (j3 + 3) = (\0\\<^sub>N, 1)" "tps ! (j3 + 4) = (\0\\<^sub>N, 1)" "tps ! (j3 + 5) = (\0\\<^sub>N, 1)" "tps ! (j3 + 6) = (\0\\<^sub>N, 1)" "tps ! (j3 + 7) = (\0\\<^sub>N, 1)" assumes "tps' = tps [j1 := nlltape (formula_n \), j3 := (\(\v. v \ set vars) \ \\\<^sub>B, 1)]" assumes "ttt = length (formula_n \) * (53 + 83 * (nlllength (formula_n \))\<^sup>2 + 67 * nlllength (formula_n \) * (nllength vars)\<^sup>2) + 13" shows "transforms (tm_sat_formula j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_sat_formula j1 j2 j3 . show ?thesis using assms loc.tps2_def loc.tm2' loc.tm2_eq_tm_sat_formula by metis qed subsection \A Turing machine for verifying \SAT{} instances\ text \ The previous Turing machine, @{const tm_sat_formula}, expects a well-formed formula and a well-formed list representing an assignment on its tapes. The TM we ultimately need, however, is not guaranteed to be given anything well-formed as input and even the well-formed inputs require decoding from the binary alphabet to the quaternary alphabet used for lists of lists of numbers. The next TM takes care of all of that and, if everything was well-formed, runs @{const tm_sat_formula}. If the first element of the pair input is invalid, it outputs \textbf{1}, as required by the definition of \SAT{}. Thus, the next Turing machine implements the function @{const verify_sat} and therefore is a verifier for \SAT. \ definition tm_verify_sat :: machine where "tm_verify_sat \ tm_right_many {0..<22} ;; tm_bindecode 0 2 ;; tm_unpair 2 3 4 ;; tm_even_length 3 5 ;; tm_proper_symbols_lt 3 6 4 ;; tm_and 6 5 ;; IF \rs. rs ! 6 \ \ THEN tm_bindecode 3 7 ;; tm_numlistlist_wf 7 8 ;; IF \rs. rs ! 8 \ \ THEN tm_proper_symbols_lt 4 10 4 ;; IF \rs. rs ! 10 \ \ THEN tm_bindecode 4 11 ;; tm_rstrip \ 11 ;; tm_numlist_wf 11 12 ;; IF \rs. rs ! 12 \ \ THEN tm_sat_formula 7 11 14 ;; tm_copyn 14 1 ELSE [] ENDIF ELSE [] ENDIF ELSE tm_setn 1 1 ENDIF ELSE tm_setn 1 1 ENDIF" lemma tm_verify_sat_tm: "turing_machine 22 6 tm_verify_sat" unfolding tm_verify_sat_def using tm_copyn_tm tm_setn_tm turing_machine_branch_turing_machine tm_sat_formula_tm tm_bindecode_tm tm_rstrip_tm tm_numlist_wf_tm tm_proper_symbols_lt_tm tm_numlistlist_wf_tm Nil_tm tm_right_many_tm tm_unpair_tm tm_even_length_tm tm_and_tm by simp locale turing_machine_verify_sat begin definition "tm1 \ tm_right_many {0..<22}" definition "tm2 \ tm1 ;; tm_bindecode 0 2" definition "tm3 \ tm2 ;; tm_unpair 2 3 4" definition "tm4 \ tm3 ;; tm_even_length 3 5" definition "tm5 \ tm4 ;; tm_proper_symbols_lt 3 6 4" definition "tm6 \ tm5 ;; tm_and 6 5" definition "tmTTT1 \ tm_bindecode 4 11" definition "tmTTT2 \ tmTTT1 ;; tm_rstrip \ 11" definition "tmTTT3 \ tmTTT2 ;; tm_numlist_wf 11 12" definition "tmTTTT1 \ tm_sat_formula 7 11 14" definition "tmTTTT2 \ tmTTTT1 ;; tm_copyn 14 1" definition "tmTTTI \ IF \rs. rs ! 12 \ \ THEN tmTTTT2 ELSE [] ENDIF" definition "tmTTT \ tmTTT3 ;; tmTTTI" definition "tmTTI \ IF \rs. rs ! 10 \ \ THEN tmTTT ELSE [] ENDIF" definition "tmTT1 \ tm_proper_symbols_lt 4 10 4" definition "tmTT \ tmTT1 ;; tmTTI" definition "tmTI \ IF \rs. rs ! 8 \ \ THEN tmTT ELSE tm_setn 1 1 ENDIF" definition "tmT1 \ tm_bindecode 3 7" definition "tmT2 \ tmT1 ;; tm_numlistlist_wf 7 8" definition "tmT \ tmT2 ;; tmTI" definition "tmI \ IF \rs. rs ! 6 \ \ THEN tmT ELSE tm_setn 1 1 ENDIF" definition "tm7 \ tm6 ;; tmI" lemma tm7_eq_tm_verify_sat: "tm7 = tm_verify_sat" unfolding tm_verify_sat_def tm7_def tmI_def tmT_def tmT2_def tmTI_def tmT1_def tmTT_def tmTT1_def tmTTI_def tmTTT_def tmTTT3_def tmTTTT1_def tmTTTI_def tmTTTT2_def tmTTT3_def tmTTT2_def tmTTT1_def tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def by simp context fixes tps0 :: "tape list" and zs :: "symbol list" assumes zs: "bit_symbols zs" assumes tps0: "tps0 = snd (start_config 22 zs)" begin definition "tps1 \ map (\tp. tp |#=| 1) tps0" lemma map_upt_length: "map f xs = map (\i. f (xs ! i)) [0..zs\, 1)" "0 < j \ j < 22 \ tps1 ! j = (\[]\, 1)" "length tps1 = 22" using tps0 start_config_def tps1_def by auto lemma tm1 [transforms_intros]: "transforms tm1 tps0 1 tps1" unfolding tm1_def proof (tform tps: tps0 tps1_def) have "length tps0 = 22" using tps0 start_config_def by simp then have "map (\j. if j \ {0..<22} then tps0 ! j |+| 1 else tps0 ! j) [0..j. tps0 ! j |+| 1) [0..j. tps0 ! j |#=| 1) [0..length tps0 = 22\ start_config_pos by simp also have "... = map (\tp. tp |#=| 1) tps0" using map_upt_length[of "\tp. tp |#=| 1" tps0] by simp also have "... = tps1" using tps1_def by simp finally show "tps1 = map (\j. if j \ {0..<22} then tps0 ! j |+| 1 else tps0 ! j) [0.. tps1 [2 := (\bindecode zs\, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 8 + 3 * length zs" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: assms zs tps1 tps2_def) definition "tps3 \ tps1 [2 := (\bindecode zs\, 1), 3 := (\first (bindecode zs)\, 1), 4 := (\second (bindecode zs)\, 1)]" lemma tm3 [transforms_intros]: assumes "ttt = 21 + 3 * length zs + 6 * length (bindecode zs)" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: assms zs tps2_def tps1 tps3_def) show "proper_symbols (bindecode zs)" using zs proper_bindecode by simp show "ttt = 8 + 3 * length zs + (6 * length (bindecode zs) + 13)" using assms by simp qed definition "tps4 \ tps1 [2 := (\bindecode zs\, 1), 3 := (\first (bindecode zs)\, 1), 4 := (\second (bindecode zs)\, 1), 5 := (\even (length (first (bindecode zs)))\\<^sub>B, 1)]" lemma tm4 [transforms_intros]: assumes "ttt = 28 + 3 * length zs + 6 * length (bindecode zs) + 7 * length (first (bindecode zs))" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: assms zs tps1 tps3_def tps4_def) show "proper_symbols (first (bindecode zs))" using zs proper_bindecode first_def by simp show "tps3 ! 5 = (\0\\<^sub>N, 1)" using tps3_def canrepr_0 tps1 by simp qed definition "tps5 \ tps1 [2 := (\bindecode zs\, 1), 3 := (\first (bindecode zs)\, 1), 4 := (\second (bindecode zs)\, 1), 5 := (\even (length (first (bindecode zs)))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first (bindecode zs))\\<^sub>B, 1)]" lemma tm5 [transforms_intros]: assumes "ttt = 33 + 3 * length zs + 6 * length (bindecode zs) + 14 * length (first (bindecode zs))" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def proof (tform tps: assms zs tps1 tps4_def tps5_def) show "proper_symbols (first (bindecode zs))" using zs proper_bindecode first_def by simp qed abbreviation "ys \ bindecode zs" abbreviation "xs \ bindecode (first ys)" abbreviation "vs \ rstrip 5 (bindecode (second ys))" definition "tps6 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1)]" lemma tm6 [transforms_intros]: assumes "ttt = 36 + 3 * length zs + 6 * length (bindecode zs) + 14 * length (first (bindecode zs))" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def by (tform tps: assms zs tps1 tps5_def tps6_def) context assumes bs_even: "proper_symbols_lt 4 (first ys) \ even (length (first ys))" begin lemma bs: "bit_symbols (first ys)" using bs_even by fastforce definition "tpsT1 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\bindecode (first ys)\, 1)]" lemma tmT1 [transforms_intros]: assumes "ttt = 7 + 3 * length (first ys)" shows "transforms tmT1 tps6 ttt tpsT1" unfolding tmT1_def by (tform tps: assms bs tps1 tps6_def tpsT1_def) definition "tpsT2 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\bindecode (first ys)\, 1), 8 := (\numlistlist_wf (bindecode (first ys))\\<^sub>B, 1)]" lemma tmT2 [transforms_intros]: assumes "ttt = 213 + 3 * length (first ys) + 39 * length (bindecode (first ys))" shows "transforms tmT2 tps6 ttt tpsT2" unfolding tmT2_def proof (tform tps: assms tps1 tpsT1_def tpsT2_def) show "proper_symbols (bindecode (first ys))" using proper_bindecode by simp show "ttt = 7 + 3 * length (first ys) + (206 + 39 * length (bindecode (first ys)))" using assms by simp qed context assumes first_wf: "numlistlist_wf (bindecode (first ys))" begin definition "tpsTT1 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\bindecode (first ys)\, 1), 8 := (\numlistlist_wf (bindecode (first ys))\\<^sub>B, 1), 10 := (\proper_symbols_lt 4 (second ys)\\<^sub>B, 1)]" lemma tmTT1 [transforms_intros]: assumes "ttt = 5 + 7 * length (second ys)" shows "transforms tmTT1 tpsT2 ttt tpsTT1" unfolding tmTT1_def proof (tform tps: tps1 tpsT2_def tpsTT1_def assms) show "proper_symbols (second ys)" using proper_bindecode second_def zs by simp qed context assumes proper_second: "proper_symbols_lt 4 (second ys)" begin definition "tpsTTT1 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\xs\, 1), 8 := (\numlistlist_wf xs\\<^sub>B, 1), 10 := (\proper_symbols_lt 4 (second ys)\\<^sub>B, 1), 11 := (\bindecode (second ys)\, 1)]" lemma tmTTT1 [transforms_intros]: assumes "ttt = 7 + 3 * length (second ys)" shows "transforms tmTTT1 tpsTT1 ttt tpsTTT1" unfolding tmTTT1_def proof (tform tps: assms tps1 tpsT2_def tpsTT1_def tpsTTT1_def) show "bit_symbols (second ys)" using proper_second by fastforce qed definition "tpsTTT2 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\xs\, 1), 8 := (\numlistlist_wf xs\\<^sub>B, 1), 10 := (\proper_symbols_lt 4 (second ys)\\<^sub>B, 1), 11 := (\vs\, 1)]" lemma tmTTT2 [transforms_intros]: assumes "ttt = 12 + 3 * length (second ys) + 3 * length (bindecode (second ys))" shows "transforms tmTTT2 tpsTT1 ttt tpsTTT2" unfolding tmTTT2_def proof (tform tps: assms tps1 tpsTTT1_def tpsTTT2_def) show "proper_symbols (bindecode (second ys))" using proper_bindecode by simp show "ttt = 7 + 3 * length (second ys) + (3 * length (bindecode (second ys)) + 5)" using assms by simp qed definition "tpsTTT3 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\xs\, 1), 8 := (\numlistlist_wf xs\\<^sub>B, 1), 10 := (\proper_symbols_lt 4 (second ys)\\<^sub>B, 1), 11 := (\vs\, 1), 12 := (\numlist_wf vs\\<^sub>B, 1)]" lemma tmTTT3 [transforms_intros]: assumes "ttt = 106 + 3 * length (second ys) + 3 * length (bindecode (second ys)) + 19 * length vs" shows "transforms tmTTT3 tpsTT1 ttt tpsTTT3" unfolding tmTTT3_def proof (tform tps: assms tps1 tpsTTT2_def tpsTTT3_def) show "proper_symbols vs" using proper_bindecode rstrip_def by simp qed context assumes second_wf: "numlist_wf vs" begin definition "tpsTTTT1 \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\xs\, 1), 8 := (\numlistlist_wf xs\\<^sub>B, 1), 10 := (\proper_symbols_lt 4 (second ys)\\<^sub>B, 1), 11 := (\vs\, 1), 12 := (\numlist_wf vs\\<^sub>B, 1), 7 := nlltape (formula_n (zs_formula xs)), 14 := (\(\v. v \ set (zs_numlist vs)) \ zs_formula xs\\<^sub>B, 1)]" lemma tmTTTT1 [transforms_intros]: assumes "ttt = length (formula_n (zs_formula xs)) * (53 + 83 * (nlllength (formula_n (zs_formula xs)))\<^sup>2 + 67 * nlllength (formula_n (zs_formula xs)) * (nllength (zs_numlist vs))\<^sup>2) + 13" shows "transforms tmTTTT1 tpsTTT3 ttt tpsTTTT1" unfolding tmTTTT1_def proof (tform time: assms) show "tpsTTT3 ! 14 = (\0\\<^sub>N, 1)" "tpsTTT3 ! (14 + 2) = (\0\\<^sub>N, 1)" "tpsTTT3 ! (14 + 3) = (\0\\<^sub>N, 1)" "tpsTTT3 ! (14 + 4) = (\0\\<^sub>N, 1)" "tpsTTT3 ! (14 + 5) = (\0\\<^sub>N, 1)" "tpsTTT3 ! (14 + 6) = (\0\\<^sub>N, 1)" "tpsTTT3 ! (14 + 7) = (\0\\<^sub>N, 1)" unfolding tpsTTT3_def using tps1 canrepr_0 by auto show "tpsTTT3 ! (14 + 1) = (\[]\\<^sub>N\<^sub>L, 1)" unfolding tpsTTT3_def using tps1 nlcontents_Nil by simp show "14 + 7 < length tpsTTT3" unfolding tpsTTT3_def using tps1 by simp let ?phi = "zs_formula xs" have "numlistlist (formula_n ?phi) = xs" using formula_zs_def formula_zs_formula first_wf by simp then have "nlltape' (formula_n ?phi) 0 = (\xs\, 1)" by (simp add: nllcontents_def) then show "tpsTTT3 ! 7 = nlltape' (formula_n ?phi) 0" unfolding tpsTTT3_def using tps1 by simp let ?vars = "zs_numlist vs" have "numlist ?vars = vs" using numlist_zs_numlist second_wf by simp then have "nltape' ?vars 0 = (\vs\, 1)" by (simp add: nlcontents_def) then show "tpsTTT3 ! 11 = nltape' ?vars 0" unfolding tpsTTT3_def using tps1 by simp show "tpsTTTT1 = tpsTTT3 [7 := nlltape (formula_n (zs_formula xs)), 14 := (\(\v. v \ set (zs_numlist vs)) \ zs_formula xs\\<^sub>B, 1)]" unfolding tpsTTTT1_def tpsTTT3_def by fast qed definition "tpsTTTT2 \ tps1 [1 := (\(\v. v \ set (zs_numlist vs)) \ zs_formula xs\\<^sub>B, 1), 2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\xs\, 1), 8 := (\numlistlist_wf xs\\<^sub>B, 1), 10 := (\proper_symbols_lt 4 (second ys)\\<^sub>B, 1), 11 := (\vs\, 1), 12 := (\numlist_wf vs\\<^sub>B, 1), 7 := nlltape (formula_n (zs_formula xs)), 14 := (\(\v. v \ set (zs_numlist vs)) \ zs_formula xs\\<^sub>B, 1)]" lemma tmTTTT2: assumes "ttt = length (formula_n (zs_formula xs)) * (53 + 83 * (nlllength (formula_n (zs_formula xs)))\<^sup>2 + 67 * nlllength (formula_n (zs_formula xs)) * (nllength (zs_numlist vs))\<^sup>2) + 27 + 3 * (nlength (if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then 1 else 0))" shows "transforms tmTTTT2 tpsTTT3 ttt tpsTTTT2" unfolding tmTTTT2_def proof (tform) show "14 < length tpsTTTT1" "1 < length tpsTTTT1" unfolding tpsTTTT1_def using tps1 by simp_all show "tpsTTTT1 ! 1 = (\0\\<^sub>N, 1)" unfolding tpsTTTT1_def using tps1 canrepr_0 by auto let ?b = "if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then 1 else 0 :: nat" show "tpsTTTT1 ! 14 = (\?b\\<^sub>N, 1)" unfolding tpsTTTT1_def using tps1 by simp show "ttt = length (formula_n (zs_formula xs)) * (53 + 83 * (nlllength (formula_n (zs_formula xs)))\<^sup>2 + 67 * nlllength (formula_n (zs_formula xs)) * (nllength (zs_numlist vs))\<^sup>2) + 13 + (14 + 3 * (nlength (if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then 1 else 0) + nlength 0))" using assms by simp show "tpsTTTT2 = tpsTTTT1 [1 := (\(\v. v \ set (zs_numlist vs)) \ zs_formula xs\\<^sub>B, 1)]" unfolding tpsTTTT2_def tpsTTTT1_def by (simp add: list_update_swap) qed lemma tmTTTT2' [transforms_intros]: assumes "ttt = 203 * length zs ^ 4 + 30" shows "transforms tmTTTT2 tpsTTT3 ttt tpsTTTT2" proof - let ?phi = "zs_formula xs" let ?ttt = "length (formula_n ?phi) * (53 + 83 * (nlllength (formula_n ?phi))\<^sup>2 + 67 * nlllength (formula_n ?phi) * (nllength (zs_numlist vs))\<^sup>2) + 27 + 3 * (nlength (if (\v. v \ set (zs_numlist vs)) \ ?phi then 1 else 0))" have "nlllength (formula_n ?phi) \ length xs" using formula_zs_def formula_zs_formula first_wf nlllength_def by simp then have 1: "nlllength (formula_n ?phi) \ length zs" by (metis div_le_dividend le_trans length_bindecode length_first) moreover have "length (formula_n ?phi) \ nlllength (formula_n ?phi)" by (simp add: length_le_nlllength) ultimately have 2: "length (formula_n ?phi) \ length zs" by simp have "nllength (zs_numlist vs) \ length vs" using second_wf numlist_zs_numlist nllength_def by simp moreover have "length vs \ length zs" using second_def length_bindecode length_rstrip_le by (metis div_le_dividend dual_order.trans length_second) ultimately have 3: "nllength (zs_numlist vs) \ length zs" by simp have 4: "nlength (if (\v. v \ set (zs_numlist vs)) \ ?phi then 1 else 0) \ 1" using nlength_1_simp by simp have "?ttt \ length (formula_n ?phi) * (53 + 83 * (nlllength (formula_n ?phi))\<^sup>2 + 67 * nlllength (formula_n ?phi) * (nllength (zs_numlist vs))\<^sup>2) + 30" using 4 by simp also have "... \ length zs * (53 + 83 * (nlllength (formula_n ?phi))\<^sup>2 + 67 * nlllength (formula_n ?phi) * (nllength (zs_numlist vs))\<^sup>2) + 30" using 2 by simp also have "... \ length zs * (53 + 83 * (length zs)\<^sup>2 + 67 * length zs * (nllength (zs_numlist vs))\<^sup>2) + 30" using 1 by (simp add: add_mono) also have "... \ length zs * (53 + 83 * (length zs)\<^sup>2 + 67 * length zs * (length zs)\<^sup>2) + 30" using 3 by simp also have "... = 53 * length zs + 83 * length zs ^ 3 + 67 * length zs ^ 4 + 30" by algebra also have "... \ 53 * length zs + 83 * length zs ^ 4 + 67 * length zs ^ 4 + 30" using pow_mono' by simp also have "... \ 53 * length zs ^ 4 + 83 * length zs ^ 4 + 67 * length zs ^ 4 + 30" using linear_le_pow by simp also have "... = 203 * length zs ^ 4 + 30" by simp finally have "?ttt \ 203 * length zs ^ 4 + 30" . then show ?thesis using assms tmTTTT2 transforms_monotone by simp qed end (* context second_wf: "numlist_wf vs" *) definition "tpsTTT \ (if numlist_wf vs then tpsTTTT2 else tpsTTT3)" lemma length_tpsTTT: "length tpsTTT = 22" using tpsTTT_def tpsTTTT2_def tpsTTT3_def tps1 by (metis (no_types, lifting) length_list_update) lemma tpsTTT: "tpsTTT ! 1 = (\if numlist_wf vs then (if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then 1 else 0) else 0\\<^sub>N, 1)" proof (cases "numlist_wf vs") case True then have "tpsTTT ! 1 = tpsTTTT2 ! 1" using tpsTTT_def by simp also have "... = (\(\v. v \ set (zs_numlist vs)) \ zs_formula xs\\<^sub>B, 1)" unfolding tpsTTTT2_def[OF True] using tps1 by simp finally show ?thesis using True by simp next case False then have "tpsTTT ! 1 = tpsTTT3 ! 1" using tpsTTT_def by simp also have "... = (\0\\<^sub>N, 1)" unfolding tpsTTT3_def using tps1 canrepr_0 by simp finally show ?thesis using False by simp qed lemma tmTTTI [transforms_intros]: assumes "ttt = 203 * length zs ^ 4 + 32" shows "transforms tmTTTI tpsTTT3 ttt tpsTTT" unfolding tmTTTI_def proof (tform time: assms) have *: "read tpsTTT3 ! 12 \ \ \ numlist_wf vs" using tpsTTT3_def tps1 read_ncontents_eq_0 by simp show "read tpsTTT3 ! 12 \ \ \ numlist_wf vs" using * by simp show "read tpsTTT3 ! 12 \ \ \ tpsTTT = tpsTTTT2" using * tpsTTT_def by simp show "\ read tpsTTT3 ! 12 \ \ \ tpsTTT = tpsTTT3" using * tpsTTT_def by simp qed lemma tmTTT: assumes "ttt = 138 + 3 * length (second ys) + 3 * length (bindecode (second ys)) + 19 * length vs + 203 * length zs ^ 4" shows "transforms tmTTT tpsTT1 ttt tpsTTT" unfolding tmTTT_def by (tform tps: assms) lemma tmTTT' [transforms_intros]: assumes "ttt = 138 + 228 * length zs ^ 4" shows "transforms tmTTT tpsTT1 ttt tpsTTT" proof - let ?ttt = "138 + 3 * length (second ys) + 3 * length (bindecode (second ys)) + 19 * length vs + 203 * length zs ^ 4" have "length ys \ length zs" by simp then have 1: "length (second ys) \ length zs" using length_second dual_order.trans by blast then have 2: "length (bindecode (second ys)) \ length zs" by simp then have 3: "length vs \ length zs" by (meson dual_order.trans length_rstrip_le) have "?ttt \ 138 + 3 * length zs + 3 * length zs + 19 * length zs + 203 * length zs ^ 4" using 1 2 3 by simp also have "... = 138 + 25 * length zs + 203 * length zs ^ 4" by simp also have "... \ 138 + 25 * length zs ^ 4 + 203 * length zs ^ 4" using linear_le_pow by simp also have "... = 138 + 228 * length zs ^ 4" by simp finally have "?ttt \ 138 + 228 * length zs ^ 4" . then show ?thesis using assms tmTTT transforms_monotone by blast qed end (* context proper_second: "proper_symbols_lt 4 (second ys)" *) definition "tpsTT \ (if proper_symbols_lt 4 (second ys) then tpsTTT else tpsTT1)" lemma length_tpsTT: "length tpsTT = 22" using tpsTT_def length_tpsTTT tpsTT1_def tps1 by simp lemma tpsTT: "tpsTT ! 1 = (ncontents (if proper_symbols_lt 4 (second ys) \ numlist_wf vs then if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then 1 else 0 else 0), 1)" proof (cases "proper_symbols_lt 4 (second ys)") case True then have "tpsTT ! 1 = tpsTTT ! 1" using tpsTT_def by simp then show ?thesis using tpsTTT True by simp next case False then have "tpsTT ! 1 = tpsTT1 ! 1" using tpsTT_def by auto then show ?thesis using tpsTT1_def tps1 canrepr_0 False by auto qed lemma tmTTI [transforms_intros]: assumes "ttt = 140 + 228 * length zs ^ 4" shows "transforms tmTTI tpsTT1 ttt tpsTT" unfolding tmTTI_def proof (tform time: assms) have *: "read tpsTT1 ! 10 \ \ \ proper_symbols_lt 4 (second ys)" using tpsTT1_def tps1 read_ncontents_eq_0 by simp show "read tpsTT1 ! 10 \ \ \ proper_symbols_lt 4 (second ys)" using * by simp let ?t = "138 + 228 * length zs ^ 4" show "read tpsTT1 ! 10 \ \ \ tpsTT = tpsTTT" using * tpsTT_def by simp show "\ read tpsTT1 ! 10 \ \ \ tpsTT = tpsTT1" using * tpsTT_def by auto qed lemma tmTT [transforms_intros]: assumes "ttt = 145 + 7 * length (second ys) + 228 * length zs ^ 4" shows "transforms tmTT tpsT2 ttt tpsTT" unfolding tmTT_def by (tform time: assms) end (* context first_wf: "numlistlist_wf (bindecode (first ys))" *) definition "tpsTE \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 7 := (\bindecode (first ys)\, 1), 8 := (\numlistlist_wf xs\\<^sub>B, 1), 1 := (\1\\<^sub>N, 1)]" definition "tpsT \ (if numlistlist_wf xs then tpsTT else tpsTE)" lemma length_tpsT: "length tpsT = 22" using tpsT_def length_tpsTT tpsTE_def tps1 by simp lemma tpsT: "tpsT ! 1 = (ncontents (if numlistlist_wf xs then if proper_symbols_lt 4 (second ys) \ numlist_wf vs then if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then 1 else 0 else 0 else 1), 1)" proof (cases "numlistlist_wf xs") case True then have "tpsT ! 1 = tpsTT ! 1" using tpsT_def by simp then show ?thesis using tpsTT True by simp next case False then have "tpsT ! 1 = tpsTE ! 1" using tpsT_def by auto then show ?thesis using tpsTE_def tps1 canrepr_0 False by auto qed lemma tmTI [transforms_intros]: assumes "ttt = 147 + 7 * length (second ys) + 228 * length zs ^ 4" shows "transforms tmTI tpsT2 ttt tpsT" unfolding tmTI_def proof (tform time: assms) have *: "read tpsT2 ! 8 \ \ \ numlistlist_wf xs" using tpsT2_def tps1 read_ncontents_eq_0 by simp show "read tpsT2 ! 8 \ \ \ numlistlist_wf xs" using * by simp show "1 < length tpsT2" using tpsT2_def tps1 by simp show "tpsT2 ! 1 = (\0\\<^sub>N, 1)" using tpsT2_def tps1 canrepr_0 by simp show "\ read tpsT2 ! 8 \ \ \ tpsT = tpsT2[1 := (\1\\<^sub>N, 1)]" using tpsT_def * tpsT2_def tpsTE_def by presburger show "read tpsT2 ! 8 \ \ \ tpsT = tpsTT" using * tpsT_def by simp show "10 + 2 * nlength 0 + 2 * nlength 1 + 1 \ ttt" using assms nlength_1_simp by simp qed lemma tmT [transforms_intros]: assumes "ttt = 360 + 3 * length (first ys) + 39 * length xs + 7 * length (second ys) + 228 * length zs ^ 4" shows "transforms tmT tps6 ttt tpsT" unfolding tmT_def by (tform time: assms) end (* context bs_even: "proper_symbols_lt 4 (first ys) \ even (length (first ys))" *) definition "tpsE \ tps1 [2 := (\ys\, 1), 3 := (\first ys\, 1), 4 := (\second ys\, 1), 5 := (\even (length (first ys))\\<^sub>B, 1), 6 := (\proper_symbols_lt 4 (first ys) \ even (length (first ys))\\<^sub>B, 1), 1 := (\1\\<^sub>N, 1)]" definition "tps7 \ (if proper_symbols_lt 4 (first ys) \ even (length (first ys)) then tpsT else tpsE)" lemma length_tps7: "length tps7 = 22" using tps7_def length_tpsT tpsE_def tps1 by simp lemma tps7: "tps7 ! 1 = (ncontents (if proper_symbols_lt 4 (first ys) \ even (length (first ys)) \ numlistlist_wf xs then if proper_symbols_lt 4 (second ys) \ numlist_wf vs then if (\v. v \ set (zs_numlist vs)) \ zs_formula xs then 1 else 0 else 0 else 1), 1)" proof (cases "proper_symbols_lt 4 (first ys) \ even (length (first ys))") case True then have "tps7 ! 1 = tpsT ! 1" using tps7_def by simp then show ?thesis using tpsT True by simp next case False then have "tps7 ! 1 = tpsE ! 1" using tps7_def by auto then show ?thesis using tpsE_def tps1 canrepr_0 False by auto qed lemma tps7': "tps7 ! 1 = (\verify_sat zs\, 1)" proof - have "proper_symbols_lt 4 zs = bit_symbols zs" for zs by fastforce then show ?thesis unfolding verify_sat_def Let_def using tps7 canrepr_0 canrepr_1 by auto qed lemma tmI [transforms_intros]: assumes "ttt = 362 + 3 * length (first ys) + 39 * length xs + 7 * length (second ys) + 228 * length zs ^ 4" shows "transforms tmI tps6 ttt tps7" unfolding tmI_def proof (tform time: assms) have *: "read tps6 ! 6 \ \ \ (proper_symbols_lt 4 (first ys)) \ even (length (first ys))" using tps6_def tps1 read_ncontents_eq_0 by simp show "read tps6 ! 6 \ \ \ (proper_symbols_lt 4 (first ys)) \ even (length (first ys))" using * by simp show "1 < length tps6" using tps6_def tps1 by simp show "tps6 ! 1 = (\0\\<^sub>N, 1)" using tps6_def tps1 canrepr_0 by simp show "\ read tps6 ! 6 \ \ \ tps7 = tps6[1 := (\1\\<^sub>N, 1)]" using tps7_def * tps6_def tpsE_def by metis show "read tps6 ! 6 \ \ \ tps7 = tpsT" using tps7_def * by simp show "10 + 2 * nlength 0 + 2 * nlength 1 + 1 \ ttt" using assms nlength_1_simp by simp qed lemma tm7: assumes "ttt = 398 + 3 * length zs + 6 * length ys + 17 * length (first ys) + 39 * length xs + 7 * length (second ys) + 228 * length zs ^ 4" shows "transforms tm7 tps0 ttt tps7" unfolding tm7_def by (tform time: assms) lemma tm7' [transforms_intros]: assumes "ttt = 398 + 300 * length zs ^ 4" shows "transforms tm7 tps0 ttt tps7" proof - have *: "length ys \ length zs" by simp then have 1: "length (second ys) \ length zs" using length_second dual_order.trans by blast have 2: "length (first ys) \ length zs" using * dual_order.trans length_first by blast then have 3: "length xs \ length zs" by simp let ?ttt = "398 + 3 * length zs + 6 * length ys + 17 * length (first ys) + 39 * length xs + 7 * length (second ys) + 228 * length zs ^ 4" have "?ttt \ 398 + 9 * length zs + 17 * length zs + 39 * length zs + 7 * length zs + 228 * length zs ^ 4" using * 1 2 3 by simp also have "... = 398 + 72 * length zs + 228 * length zs ^ 4" by simp also have "... \ 398 + 72 * length zs ^ 4 + 228 * length zs ^ 4" using linear_le_pow by simp also have "... = 398 + 300 * length zs ^ 4" by simp finally have "?ttt \ 398 + 300 * length zs ^ 4" . then show ?thesis using assms tm7 transforms_monotone by fast qed end (* context *) end (* locale *) lemma transforms_tm_verify_sat: fixes zs :: "symbol list" and tps :: "tape list" assumes "bit_symbols zs" and "tps = snd (start_config 22 zs)" and "ttt = 398 + 300 * length zs ^ 4" shows "\tps'. tps' ! 1 = (\verify_sat zs\, 1) \ transforms tm_verify_sat tps ttt tps'" proof - interpret loc: turing_machine_verify_sat . show ?thesis using assms loc.tm7' loc.tps7' loc.tm7_eq_tm_verify_sat by metis qed text \ With the Turing machine just constructed and the polynomial $p(n) = n$ we can satisfy the definition of $\NP$ and prove the main result of this chapter. \ theorem SAT_in_NP: "SAT \ \\

" proof - define p :: "nat \ nat" where "p = (\n. n)" define T :: "nat \ nat" where "T = (\n. 398 + 300 * n ^ 4)" define f :: "string \ string" where "f = (\x. symbols_to_string (verify_sat (string_to_symbols x)))" have "turing_machine 22 6 tm_verify_sat" using tm_verify_sat_tm . moreover have "polynomial p" using p_def polynomial_id by (metis eq_id_iff) moreover have "big_oh_poly T" using T_def big_oh_poly_poly big_oh_poly_const big_oh_poly_sum big_oh_poly_prod by simp moreover have "computes_in_time 22 tm_verify_sat f T" proof fix x :: string let ?zs = "string_to_symbols x" have bs: "bit_symbols ?zs" by simp have "bit_symbols (verify_sat ?zs)" using bit_symbols_verify_sat by simp then have *: "string_to_symbols (f x) = verify_sat ?zs" unfolding f_def using bit_symbols_to_symbols by simp obtain tps where tps: "tps ! 1 = (\verify_sat ?zs\, 1)" "transforms tm_verify_sat (snd (start_config 22 ?zs)) (T (length ?zs)) tps" using bs transforms_tm_verify_sat T_def by blast then have "tps ::: 1 = string_to_contents (f x)" using * start_config_def contents_string_to_contents by simp then show "\tps. tps ::: 1 = string_to_contents (f x) \ transforms tm_verify_sat (snd (start_config_string 22 x)) (T (length x)) tps" using tps(2) by auto qed moreover have "\x. x \ SAT \ (\u. length u = p (length x) \ f \x, u\ = [\])" proof fix x :: string show "(x \ SAT) = (\u. length u = p (length x) \ f \x, u\ = [\])" proof show "\u. length u = p (length x) \ f \x, u\ = [\]" if "x \ SAT" proof (cases "\\. x = formula_to_string \") case True then obtain \ where \: "x = formula_to_string \" "satisfiable \" using SAT_def using \x \ SAT\ by auto then obtain us where us: "bit_symbols us" "length us = length (formula_to_string \)" "verify_sat \formula_to_string \; symbols_to_string us\ = [\]" using ex_witness_linear_length by blast let ?zs = "\formula_to_string \; symbols_to_string us\" define u where "u = symbols_to_string us" have "length us = p (length x)" using us(2) \(1) p_def by simp then have 1: "length u = p (length x)" using u_def by simp have "f \x, u\ = symbols_to_string (verify_sat \x; u\)" using f_def by simp also have "... = symbols_to_string (verify_sat ?zs)" using \(1) u_def by simp also have "... = symbols_to_string [\]" using us(3) by simp also have "... = [\]" by simp finally have "f \x, u\ = [\]" . then show ?thesis using 1 by auto next case False define u where "u = replicate (length x) \" then have 1: "length u = p (length x)" using p_def by simp have "f \x, u\ = symbols_to_string (verify_sat \x; u\)" using f_def by simp also have "... = symbols_to_string [\]" using verify_sat_not_wf_phi False by simp also have "... = [\]" by simp finally have "f \x, u\ = [\]" . then show ?thesis using 1 by auto qed show "x \ SAT" if ex: "\u. length u = p (length x) \ f \x, u\ = [\]" proof (rule ccontr) assume notin: "x \ SAT" then obtain \ where \: "x = formula_to_string \" "\ satisfiable \" using SAT_def by auto obtain u where u: "length u = p (length x)" "f \x, u\ = [\]" using ex by auto have "f \x, u\ = symbols_to_string (verify_sat \x; u\)" using f_def by simp also have "... = symbols_to_string (verify_sat \formula_to_string \; u\)" using \(1) by simp also have "... = symbols_to_string []" using verify_sat_not_sat \(2) by simp also have "... = []" by simp finally have "f \x, u\ = []" . then show False using u(2) by simp qed qed qed ultimately show ?thesis using complexity_class_NP_def by fast qed end diff --git a/thys/Cook_Levin/Symbol_Ops.thy b/thys/Cook_Levin/Symbol_Ops.thy --- a/thys/Cook_Levin/Symbol_Ops.thy +++ b/thys/Cook_Levin/Symbol_Ops.thy @@ -1,1986 +1,1986 @@ section \Symbol sequence operations\ theory Symbol_Ops imports Two_Four_Symbols begin text \ While previous sections have focused on ``formatted'' symbol sequences for numbers and lists, in this section we devise some Turing machines dealing with ``unstructured'' arbitrary symbol sequences. The only ``structure'' that is often imposed is that of not containing a blank symbol because when reading a symbol sequence, say from the input tape, a blank would signal the end of the symbol sequence. \ subsection \Checking for being over an alphabet\ text \ In this section we devise a Turing machine that checks if a proper symbol sequence is over a given alphabet represented by an upper bound symbol $z$. \ abbreviation proper_symbols_lt :: "symbol \ symbol list \ bool" where "proper_symbols_lt z zs \ proper_symbols zs \ symbols_lt z zs" text \ The next Turing machine checks if the symbol sequence (up until the first blank) on tape $j_1$ contains only symbols from $\{2, \dots, z - 1\}$, where $z$ is a parameter of the TM, and writes to tape $j_2$ the number~1 or~0, representing True or False, respectively. It assumes that $j_2$ initially contains at most one symbol. \ definition tm_proper_symbols_lt :: "tapeidx \ tapeidx \ symbol \ machine" where "tm_proper_symbols_lt j1 j2 z \ tm_write j2 \ ;; WHILE [] ; \rs. rs ! j1 \ \ DO IF \rs. rs ! j1 < 2 \ rs ! j1 \ z THEN tm_write j2 \ ELSE [] ENDIF ;; tm_right j1 DONE ;; tm_cr j1" lemma tm_proper_symbols_lt_tm: assumes "0 < j2" "j1 < k" "j2 < k" and "G \ 4" shows "turing_machine k G (tm_proper_symbols_lt j1 j2 z)" using assms tm_write_tm tm_right_tm tm_cr_tm Nil_tm tm_proper_symbols_lt_def turing_machine_loop_turing_machine turing_machine_branch_turing_machine by simp locale turing_machine_proper_symbols_lt = fixes j1 j2 :: tapeidx and z :: symbol begin definition "tm1 \ tm_write j2 \" definition "tm2 \ IF \rs. rs ! j1 < 2 \ rs ! j1 \ z THEN tm_write j2 \ ELSE [] ENDIF" definition "tm3 \ tm2 ;; tm_right j1" definition "tm4 \ WHILE [] ; \rs. rs ! j1 \ \ DO tm3 DONE" definition "tm5 \ tm1 ;; tm4" definition "tm6 \ tm5 ;; tm_cr j1" lemma tm6_eq_tm_proper_symbols_lt: "tm6 = tm_proper_symbols_lt j1 j2 z" unfolding tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def tm_proper_symbols_lt_def by simp context fixes zs :: "symbol list" and tps0 :: "tape list" and k :: nat assumes jk: "k = length tps0" "j1 \ j2" "j1 < k" "j2 < k" and zs: "proper_symbols zs" and tps0: "tps0 ! j1 = (\zs\, 1)" "tps0 ! j2 = (\[]\, 1)" begin definition "tps1 t \ tps0 [j1 := (\zs\, Suc t), j2 := (if proper_symbols_lt z (take t zs) then \[\]\ else \[]\, 1)]" lemma tm1 [transforms_intros]: "transforms tm1 tps0 1 (tps1 0)" unfolding tm1_def proof (tform tps: jk tps0) have "(if proper_symbols_lt z (take 0 zs) then \[\]\ else \[]\, 1) = (\[\]\, 1)" by simp moreover have "tps0 ! j2 |:=| \ = (\[\]\, 1)" using tps0(2) contents_def by auto moreover have "tps0[j1 := (\zs\, Suc 0)] = tps0" using tps0(1) by (metis One_nat_def list_update_id) ultimately show "tps1 0 = tps0[j2 := tps0 ! j2 |:=| \]" unfolding tps1_def by auto qed definition "tps2 t \ tps0 [j1 := (\zs\, Suc t), j2 := (if proper_symbols_lt z (take (Suc t) zs) then \[\]\ else \[]\, 1)]" lemma tm2 [transforms_intros]: assumes "t < length zs" shows "transforms tm2 (tps1 t) 3 (tps2 t)" unfolding tm2_def proof (tform tps: jk tps1_def) have "tps1 t ! j1 = (\zs\, Suc t)" using tps1_def jk by simp moreover have "read (tps1 t) ! j1 = tps1 t :.: j1" using tapes_at_read' jk tps1_def by (metis (no_types, lifting) length_list_update) ultimately have *: "read (tps1 t) ! j1 = zs ! t" using contents_inbounds assms(1) by simp have j2: "tps1 t ! j2 = (if proper_symbols_lt z (take t zs) then \[\]\ else \[]\, 1)" using tps1_def jk by simp show "tps2 t = (tps1 t)[j2 := tps1 t ! j2 |:=| \]" if "read (tps1 t) ! j1 < 2 \ z \ read (tps1 t) ! j1" proof - have c3: "(\[\]\, 1) |:=| \ = (\[]\, 1)" using contents_def by auto have "(if proper_symbols_lt z (take t zs) then \[\]\ else \[]\, 1) |:=| \ = (if proper_symbols_lt z (take (Suc t) zs) then \[\]\ else \[]\, 1)" proof (cases "proper_symbols_lt z (take t zs)") case True have "zs ! t < 2 \ z \ zs ! t" using that * by simp then have "\ proper_symbols_lt z (take (Suc t) zs)" using assms(1) by auto then show ?thesis using c3 by auto next case False then have "\ proper_symbols_lt z (take (Suc t) zs)" by auto then show ?thesis using c3 False by auto qed then have "tps1 t ! j2 |:=| \ = (if proper_symbols_lt z (take (Suc t) zs) then \[\]\ else \[]\, 1)" using j2 by simp then show "tps2 t = (tps1 t)[j2 := tps1 t ! j2 |:=| \]" unfolding tps2_def tps1_def using c3 jk(1,4) by simp qed show "tps2 t = tps1 t" if "\ (read (tps1 t) ! j1 < 2 \ z \ read (tps1 t) ! j1)" proof - have 1: "zs ! t \ 2 \ z > zs ! t" using that * by simp show "tps2 t = tps1 t" proof (cases "proper_symbols_lt z (take t zs)") case True have "proper_symbols_lt z (take (Suc t) zs)" using True 1 assms(1) zs by (metis length_take less_antisym min_less_iff_conj nth_take) then show ?thesis using tps1_def tps2_def jk by simp next case False then have "\ proper_symbols_lt z (take (Suc t) zs)" by auto then show ?thesis using tps1_def tps2_def jk False by auto qed qed qed lemma tm3 [transforms_intros]: assumes "t < length zs" shows "transforms tm3 (tps1 t) 4 (tps1 (Suc t))" unfolding tm3_def proof (tform tps: assms jk tps2_def) have "tps2 t ! j1 |+| 1 = (\zs\, Suc (Suc t))" using tps2_def jk by simp then show "tps1 (Suc t) = (tps2 t)[j1 := tps2 t ! j1 |+| 1]" unfolding tps1_def tps2_def by (metis (no_types, lifting) jk(2) list_update_overwrite list_update_swap) qed lemma tm4 [transforms_intros]: assumes "ttt = 1 + 6 * length zs" shows "transforms tm4 (tps1 0) ttt (tps1 (length zs))" unfolding tm4_def proof (tform time: assms) show "read (tps1 t) ! j1 \ \" if "t < length zs" for t proof - have "tps1 t ! j1 = (\zs\, Suc t)" using tps1_def jk by simp moreover have "read (tps1 t) ! j1 = tps1 t :.: j1" using tapes_at_read' jk tps1_def by (metis (no_types, lifting) length_list_update) ultimately have "read (tps1 t) ! j1 = zs ! t" using contents_inbounds that by simp then show ?thesis using zs that by auto qed show "\ read (tps1 (length zs)) ! j1 \ \" proof - have "tps1 (length zs) ! j1 = (\zs\, Suc (length zs))" using tps1_def jk by simp moreover have "read (tps1 (length zs)) ! j1 = tps1 (length zs) :.: j1" using tapes_at_read' jk tps1_def by (metis (no_types, lifting) length_list_update) ultimately show ?thesis by simp qed qed lemma tm5 [transforms_intros]: assumes "ttt = 2 + 6 * length zs" shows "transforms tm5 tps0 ttt (tps1 (length zs))" unfolding tm5_def by (tform time: assms) definition "tps5 \ tps0 [j1 := (\zs\, 1), j2 := (if proper_symbols_lt z zs then \[\]\ else \[]\, 1)]" definition "tps5' \ tps0 [j2 := (if proper_symbols_lt z zs then \[\]\ else \[]\, 1)]" lemma tm6: assumes "ttt = 5 + 7 * length zs" shows "transforms tm6 tps0 ttt tps5'" unfolding tm6_def proof (tform time: assms tps1_def jk) have *: "tps1 (length zs) ! j1 = (\zs\, Suc (length zs))" using tps1_def jk by simp show "clean_tape (tps1 (length zs) ! j1)" using * zs by simp have "tps5 = (tps1 (length zs))[j1 := (\zs\, Suc (length zs)) |#=| 1]" unfolding tps5_def tps1_def by (simp add: list_update_swap[OF jk(2)]) then have "tps5 = (tps1 (length zs))[j1 := tps1 (length zs) ! j1 |#=| 1]" using * by simp moreover have "tps5' = tps5" using tps5'_def tps5_def tps0 jk by (metis list_update_id) ultimately show "tps5' = (tps1 (length zs))[j1 := tps1 (length zs) ! j1 |#=| 1]" by simp qed definition "tps6 \ tps0 [j2 := (\proper_symbols_lt z zs\\<^sub>B, 1)]" lemma tm6': assumes "ttt = 5 + 7 * length zs" shows "transforms tm6 tps0 ttt tps6" proof - have "tps6 = tps5'" using tps6_def tps5'_def canrepr_0 canrepr_1 by auto then show ?thesis using tm6 assms by simp qed end end (* locale *) lemma transforms_tm_proper_symbols_ltI [transforms_intros]: fixes j1 j2 :: tapeidx and z :: symbol fixes zs :: "symbol list" and tps tps' :: "tape list" and k :: nat assumes "k = length tps" "j1 \ j2" "j1 < k" "j2 < k" and "proper_symbols zs" assumes "tps ! j1 = (\zs\, 1)" "tps ! j2 = (\[]\, 1)" assumes "ttt = 5 + 7 * length zs" assumes "tps' = tps [j2 := (\proper_symbols_lt z zs\\<^sub>B, 1)]" shows "transforms (tm_proper_symbols_lt j1 j2 z) tps ttt tps'" proof - interpret loc: turing_machine_proper_symbols_lt j1 j2 . show ?thesis using assms loc.tm6_eq_tm_proper_symbols_lt loc.tps6_def loc.tm6' by simp qed subsection \The length of the input\ text \ The Turing machine in this section reads the input tape until the first blank and increments a counter on tape $j$ for every symbol read. In the end it performs a carriage return on the input tape, and tape $j$ contains the length of the input as binary number. For this to work, tape $j$ must initially be empty. \ lemma proper_tape_read: assumes "proper_symbols zs" shows "|.| (\zs\, i) = \ \ i > length zs" proof - have "|.| (\zs\, i) = \" if "i > length zs" for i using that contents_outofbounds by simp moreover have "|.| (\zs\, i) \ \" if "i \ length zs" for i using that contents_inbounds assms contents_def proper_symbols_ne0 by simp ultimately show ?thesis by (meson le_less_linear) qed definition tm_length_input :: "tapeidx \ machine" where "tm_length_input j \ WHILE [] ; \rs. rs ! 0 \ \ DO tm_incr j ;; tm_right 0 DONE ;; tm_cr 0" lemma tm_length_input_tm: assumes "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_length_input j)" using tm_length_input_def tm_incr_tm assms Nil_tm tm_right_tm tm_cr_tm by (simp add: turing_machine_loop_turing_machine) locale turing_machine_length_input = fixes j :: tapeidx begin definition "tmL1 \ tm_incr j" definition "tmL2 \ tmL1 ;; tm_right 0" definition "tm1 \ WHILE [] ; \rs. rs ! 0 \ \ DO tmL2 DONE" definition "tm2 \ tm1 ;; tm_cr 0" lemma tm2_eq_tm_length_input: "tm2 = tm_length_input j" unfolding tm2_def tm1_def tmL2_def tmL1_def tm_length_input_def by simp context fixes tps0 :: "tape list" and k :: nat and zs :: "symbol list" assumes jk: "0 < j" "j < k" "length tps0 = k" and zs: "proper_symbols zs" and tps0: "tps0 ! 0 = (\zs\, 1)" "tps0 ! j = (\0\\<^sub>N, 1)" begin definition tpsL :: "nat \ tape list" where "tpsL t \ tps0[0 := (\zs\, 1 + t), j := (\t\\<^sub>N, 1)]" lemma tpsL_eq_tps0: "tpsL 0 = tps0" using tpsL_def tps0 jk by (metis One_nat_def list_update_id plus_1_eq_Suc) definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0[0 := (\zs\, 1 + t), j := (\Suc t\\<^sub>N, 1)]" definition tpsL2 :: "nat \ tape list" where "tpsL2 t \ tps0[0 := (\zs\, 1 + Suc t), j := (\Suc t\\<^sub>N, 1)]" lemma tmL1 [transforms_intros]: assumes "t < length zs" and "ttt = 5 + 2 * nlength t" shows "transforms tmL1 (tpsL t) ttt (tpsL1 t)" unfolding tmL1_def by (tform tps: assms(1) tpsL_def tpsL1_def tps0 jk time: assms(2)) lemma tmL2: assumes "t < length zs" and "ttt = 6 + 2 * nlength t" shows "transforms tmL2 (tpsL t) ttt (tpsL (Suc t))" unfolding tmL2_def proof (tform tps: assms(1) tpsL_def tpsL1_def tps0 jk time: assms(2)) have "tpsL1 t ! 0 = (\zs\, 1 + t)" using tpsL2_def tpsL1_def jk tps0 by simp then have "tpsL2 t = (tpsL1 t)[0 := tpsL1 t ! 0 |#=| Suc (tpsL1 t :#: 0)]" using tpsL2_def tpsL1_def jk tps0 - by (smt (z3) fstI list_update_overwrite list_update_swap nat_neq_iff plus_1_eq_Suc prod.sel(2)) + by (smt (verit) fstI list_update_overwrite list_update_swap nat_neq_iff plus_1_eq_Suc prod.sel(2)) then show "tpsL (Suc t) = (tpsL1 t)[0 := tpsL1 t ! 0 |+| 1]" using tpsL2_def tpsL_def tpsL1_def jk tps0 by simp qed lemma tmL2': assumes "t < length zs" and "ttt = 6 + 2 * nlength (length zs)" shows "transforms tmL2 (tpsL t) ttt (tpsL (Suc t))" proof - have "6 + 2 * nlength t \ 6 + 2 * nlength (length zs)" using assms(1) nlength_mono by simp then show ?thesis using assms tmL2 transforms_monotone by blast qed lemma tm1: assumes "ttt = length zs * (8 + 2 * nlength (length zs)) + 1" shows "transforms tm1 (tpsL 0) ttt (tpsL (length zs))" unfolding tm1_def proof (tform) let ?t = "6 + 2 * nlength (length zs)" show "\t. t < length zs \ transforms tmL2 (tpsL t) ?t (tpsL (Suc t))" using tmL2' by simp have *: "tpsL t ! 0 = (\zs\, Suc t)" for t using tpsL_def jk by simp then show "\t. t < length zs \ read (tpsL t) ! 0 \ \" using proper_tape_read[OF zs] tpsL_def jk tapes_at_read' by (metis length_list_update less_Suc_eq_0_disj not_less_eq) show "\ read (tpsL (length zs)) ! 0 \ \" using proper_tape_read[OF zs] tpsL_def jk tapes_at_read' * by (metis length_list_update less_Suc_eq_0_disj less_imp_Suc_add nat_neq_iff not_less_less_Suc_eq) show "length zs * (6 + 2 * nlength (length zs) + 2) + 1 \ ttt" using assms by simp qed lemma tmL' [transforms_intros]: assumes "ttt = 10 * length zs ^ 2 + 1" shows "transforms tm1 (tpsL 0) ttt (tpsL (length zs))" proof - let ?ttt = "length zs * (8 + 2 * nlength (length zs)) + 1" have "?ttt \ length zs * (8 + 2 * length zs) + 1" using nlength_le_n by simp also have "... \ 8 * length zs + 2 * length zs ^ 2 + 1" by (simp add: add_mult_distrib2 power2_eq_square) also have "... \ 10 * length zs ^ 2 + 1" using linear_le_pow by simp finally have "?ttt \ 10 * length zs ^ 2 + 1" . then show ?thesis using tm1 assms transforms_monotone by simp qed definition tps2 :: "tape list" where "tps2 \ tps0[0 := (\zs\, 1), j := (\length zs\\<^sub>N, 1)]" lemma tm2: assumes "ttt = 10 * length zs ^ 2 + length zs + 4" shows "transforms tm2 (tpsL 0) ttt tps2" unfolding tm2_def proof (tform time: assms tpsL_def jk tps: tpsL_def tpsL1_def tps0 jk) show "clean_tape (tpsL (length zs) ! 0)" using tpsL_def jk clean_contents_proper[OF zs] by simp have "tpsL (length zs) ! 0 = (\zs\, Suc (length zs))" using tpsL_def jk by simp then show "tps2 = (tpsL (length zs))[0 := tpsL (length zs) ! 0 |#=| 1]" using tps2_def tpsL_def jk by (simp add: list_update_swap_less) qed definition tps2' :: "tape list" where "tps2' \ tps0[j := (\length zs\\<^sub>N, 1)]" lemma tm2': assumes "ttt = 11 * length zs ^ 2 + 4" shows "transforms tm2 tps0 ttt tps2'" proof - have "10 * length zs ^ 2 + length zs + 4 \ ttt" using assms linear_le_pow[of 2 "length zs"] by simp moreover have "tps2 = tps2'" using tps2_def tps2'_def jk tps0 by (metis list_update_id) ultimately show ?thesis using transforms_monotone tm2 tpsL_eq_tps0 by simp qed end end lemma transforms_tm_length_inputI [transforms_intros]: fixes j :: tapeidx fixes tps tps' :: "tape list" and k :: nat and zs :: "symbol list" assumes "0 < j" "j < k" "length tps = k" and "proper_symbols zs" assumes "tps ! 0 = (\zs\, 1)" "tps ! j = (\0\\<^sub>N, 1)" assumes "ttt = 11 * length zs ^ 2 + 4" assumes "tps' = tps [j := (\length zs\\<^sub>N, 1)]" shows "transforms (tm_length_input j) tps ttt tps'" proof - interpret loc: turing_machine_length_input j . show ?thesis using loc.tm2' loc.tps2'_def loc.tm2_eq_tm_length_input assms by simp qed subsection \Whether the length is even\ text \ The next Turing machines reads the symbols on tape $j_1$ until the first blank and alternates between numbers~0 and~1 on tape $j_2$. Then tape $j_2$ contains the parity of the length of the symbol sequence on tape $j_1$. Finally, the TM flips the output once more, so that tape $j_2$ contains a Boolean indicating whether the length was even or not. We assume tape $j_2$ is initially empty, that is, represents the number~0. \ definition tm_even_length :: "tapeidx \ tapeidx \ machine" where "tm_even_length j1 j2 \ WHILE [] ; \rs. rs ! j1 \ \ DO tm_not j2 ;; tm_right j1 DONE ;; tm_not j2 ;; tm_cr j1" lemma tm_even_length_tm: assumes "k \ 2" and "G \ 4" and "j1 < k" "0 < j2" "j2 < k" shows "turing_machine k G (tm_even_length j1 j2)" using tm_even_length_def tm_right_tm tm_not_tm Nil_tm assms tm_cr_tm turing_machine_loop_turing_machine by simp locale turing_machine_even_length = fixes j1 j2 :: tapeidx begin definition "tmB \ tm_not j2 ;; tm_right j1" definition "tmL \ WHILE [] ; \rs. rs ! j1 \ \ DO tmB DONE" definition "tm2 \ tmL ;; tm_not j2" definition "tm3 \ tm2 ;; tm_cr j1" lemma tm3_eq_tm_even_length: "tm3 = tm_even_length j1 j2" unfolding tm3_def tm2_def tmL_def tmB_def tm_even_length_def by simp context fixes tps0 :: "tape list" and k :: nat and zs :: "symbol list" assumes zs: "proper_symbols zs" assumes jk: "j1 < k" "j2 < k" "j1 \ j2" "length tps0 = k" assumes tps0: "tps0 ! j1 = (\zs\, 1)" "tps0 ! j2 = (\0\\<^sub>N, 1)" begin definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j1 := (\zs\, Suc t), j2 := (\odd t\\<^sub>B, 1)]" lemma tpsL0: "tpsL 0 = tps0" unfolding tpsL_def using tps0 jk by (metis (mono_tags, opaque_lifting) One_nat_def even_zero list_update_id) lemma tmL2 [transforms_intros]: "transforms tmB (tpsL t) 4 (tpsL (Suc t))" unfolding tmB_def proof (tform tps: tpsL_def jk) have "(tpsL t) [j2 := (\(if odd t then 1 else 0 :: nat) \ 1\\<^sub>B, 1), j1 := (tpsL t)[j2 := (\ (if odd t then 1 else 0 :: nat) \ 1 \\<^sub>B, 1)] ! j1 |+| 1] = (tpsL t) [j2 := (\odd (Suc t)\\<^sub>B, 1), j1 := (tpsL t) ! j1 |+| 1]" using jk by simp also have "... = (tpsL t) [j2 := (\odd (Suc t)\\<^sub>B, 1), j1 := (\zs\, Suc (Suc t))]" using tpsL_def jk by simp also have "... = (tpsL t) [j1 := (\zs\, Suc (Suc t)), j2 := (\odd (Suc t)\\<^sub>B, 1)]" using jk by (simp add: list_update_swap) also have "... = tps0 [j1 := (\zs\, Suc (Suc t)), j2 := (\odd (Suc t)\\<^sub>B, 1)]" using jk tpsL_def by (simp add: list_update_swap) also have "... = tpsL (Suc t)" using tpsL_def by simp finally show "tpsL (Suc t) = (tpsL t) [j2 := (\(if odd t then 1 else 0 :: nat) \ 1\\<^sub>B, 1), j1 := (tpsL t)[j2 := (\(if odd t then 1 else 0 :: nat) \ 1\\<^sub>B, 1)] ! j1 |+| 1]" by simp qed lemma tmL: assumes "ttt = 6 * length zs + 1" shows "transforms tmL (tpsL 0) ttt (tpsL (length zs))" unfolding tmL_def proof (tform time: assms) have "read (tpsL t) ! j1 = tpsL t :.: j1" for t using tpsL_def tapes_at_read' jk by (metis (no_types, lifting) length_list_update) then have "read (tpsL t) ! j1 = \zs\ (Suc t)" for t using tpsL_def jk by simp then show "\t. t < length zs \ read (tpsL t) ! j1 \ \" and "\ read (tpsL (length zs)) ! j1 \ \" using zs by auto qed lemma tmL' [transforms_intros]: assumes "ttt = 6 * length zs + 1" shows "transforms tmL tps0 ttt (tpsL (length zs))" using assms tmL tpsL0 by simp definition tps2 :: "tape list" where "tps2 \ tps0 [j1 := (\zs\, Suc (length zs)), j2 := (\even (length zs) \\<^sub>B, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 6 * length zs + 4" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tpsL_def jk time: assms) show "tps2 = (tpsL (length zs))[j2 := (\(if odd (length zs) then 1 else 0 :: nat) \ 1\\<^sub>B, 1)]" unfolding tps2_def tpsL_def by (simp add: list_update_swap) qed definition tps3 :: "tape list" where "tps3 \ tps0 [j1 := (\zs\, 1), j2 := (\even (length zs)\\<^sub>B, 1)]" lemma tm3: assumes "ttt = 7 * length zs + 7" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps2_def jk time: assms) show "clean_tape (tps2 ! j1)" using tps2_def jk zs clean_contents_proper by simp have "tps2 ! j1 |#=| 1 = (\zs\, 1)" using tps2_def jk by simp then show "tps3 = tps2[j1 := tps2 ! j1 |#=| 1]" unfolding tps3_def tps2_def using jk by (simp add: list_update_swap) show "ttt = 6 * length zs + 4 + (tps2 :#: j1 + 2)" using assms tps2_def jk by simp qed definition tps3' :: "tape list" where "tps3' \ tps0 [j2 := (\even (length zs)\\<^sub>B, 1)]" lemma tps3': "tps3' = tps3" using tps3'_def tps3_def tps0 by (metis list_update_id) lemma tm3': assumes "ttt = 7 * length zs + 7" shows "transforms tm3 tps0 ttt tps3'" using tps3' tm3 assms by simp end (* context *) end (* locale *) lemma transforms_tm_even_lengthI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and k :: nat and zs :: "symbol list" assumes "j1 < k" "j2 < k" "j1 \ j2" and "proper_symbols zs" and "length tps = k" assumes "tps ! j1 = (\zs\, 1)" "tps ! j2 = (\0\\<^sub>N, 1)" assumes "tps' = tps [j2 := (\even (length zs)\\<^sub>B, 1)]" assumes "ttt = 7 * length zs + 7" shows "transforms (tm_even_length j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_even_length j1 j2 . show ?thesis using assms loc.tps3'_def loc.tm3' loc.tm3_eq_tm_even_length by simp qed subsection \Checking for ends-with or empty\ text \ The next Turing machine implements a slightly idiosyncratic operation that we use in the next section for checking if a symbol sequence represents a list of numbers. The operation consists in checking if the symbol sequence on tape $j_1$ either is empty or ends with the symbol $z$, which is another parameter of the TM. If the condition is met, the number~1 is written to tape $j_2$, otherwise the number~0. \ definition tm_empty_or_endswith :: "tapeidx \ tapeidx \ symbol \ machine" where "tm_empty_or_endswith j1 j2 z \ tm_right_until j1 {\} ;; tm_left j1 ;; IF \rs. rs ! j1 \ {\, z} THEN tm_setn j2 1 ELSE [] ENDIF ;; tm_cr j1" lemma tm_empty_or_endswith_tm: assumes "k \ 2" and "G \ 4" and "0 < j2" and "j1 < k" and "j2 < k" shows "turing_machine k G (tm_empty_or_endswith j1 j2 z)" using assms Nil_tm tm_right_until_tm tm_left_tm tm_setn_tm tm_cr_tm turing_machine_branch_turing_machine tm_empty_or_endswith_def by simp locale turing_machine_empty_or_endswith = fixes j1 j2 :: tapeidx and z :: symbol begin definition "tm1 \ tm_right_until j1 {\}" definition "tm2 \ tm1 ;; tm_left j1" definition "tmI \ IF \rs. rs ! j1 \ {\, z} THEN tm_setn j2 1 ELSE [] ENDIF" definition "tm3 \ tm2 ;; tmI" definition "tm4 \ tm3 ;; tm_cr j1" lemma tm4_eq_tm_empty_or_endswith: "tm4 = tm_empty_or_endswith j1 j2 z" unfolding tm4_def tm3_def tmI_def tm2_def tm1_def tm_empty_or_endswith_def by simp context fixes tps0 :: "tape list" and k :: nat and zs :: "symbol list" assumes jk: "j1 \ j2" "j1 < k" "j2 < k" "length tps0 = k" and zs: "proper_symbols zs" and tps0: "tps0 ! j1 = (\zs\, 1)" "tps0 ! j2 = (\0\\<^sub>N, 1)" begin definition tps1 :: "tape list" where "tps1 \ tps0 [j1 := (\zs\, Suc (length zs))]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (length zs)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform time: assms tps: tps0 tps1_def jk) show "rneigh (tps0 ! j1) {0} (length zs)" proof (rule rneighI) show "(tps0 ::: j1) (tps0 :#: j1 + length zs) \ {0}" using tps0 by simp show "\n'. n' < length zs \ (tps0 ::: j1) (tps0 :#: j1 + n') \ {0}" using zs tps0 by auto qed qed definition tps2 :: "tape list" where "tps2 \ tps0 [j1 := (\zs\, length zs)]" lemma tm2 [transforms_intros]: assumes "ttt = 2 + length zs" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform time: assms tps: tps1_def tps2_def jk) definition tps3 :: "tape list" where "tps3 \ tps0 [j1 := (\zs\, length zs), j2 := (\zs = [] \ last zs = z\\<^sub>B, 1)]" lemma tmI [transforms_intros]: "transforms tmI tps2 14 tps3" unfolding tmI_def proof (tform tps: tps0 tps2_def jk) have *: "read tps2 ! j1 = \zs\ (length zs)" using tps2_def jk tapes_at_read'[of j1 tps2] by simp show "tps3 = tps2[j2 := (\1\\<^sub>N, 1)]" if "read tps2 ! j1 \ {\, z}" proof - have "zs = [] \ last zs = z" using that * contents_inbounds zs by (metis diff_less dual_order.refl insert_iff last_conv_nth length_greater_0_conv proper_symbols_ne1 singletonD zero_less_one) then have "(if zs = [] \ last zs = z then 1 else 0) = 1" by simp then show ?thesis using tps2_def tps3_def jk by (smt (verit, best)) qed show "tps3 = tps2" if "read tps2 ! j1 \ {\, z}" proof - have "\ (zs = [] \ last zs = z)" using that * contents_inbounds zs by (metis contents_at_0 dual_order.refl insertCI last_conv_nth length_greater_0_conv list.size(3)) then have "(if zs = [] \ last zs = z then 1 else 0) = 0" by simp then show ?thesis using tps2_def tps3_def jk tps0 by (smt (verit, best) list_update_id nth_list_update_neq) qed show "10 + 2 * nlength 0 + 2 * nlength 1 + 2 \ 14" using nlength_1_simp by simp qed lemma tm3 [transforms_intros]: assumes "ttt = 16 + length zs" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def by (tform tps: assms) definition tps4 :: "tape list" where "tps4 \ tps0 [j2 := (\zs = [] \ last zs = z\\<^sub>B, 1)]" lemma tm4: assumes "ttt = 18 + 2 * length zs" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform time: assms tps3_def jk tps: tps3_def jk zs) have "tps3 ! j1 |#=| 1 = (\zs\, 1)" using tps3_def jk zs by simp then show "tps4 = tps3[j1 := tps3 ! j1 |#=| 1]" using tps4_def tps3_def jk tps0(1) by (metis list_update_id list_update_overwrite list_update_swap) qed end (* context *) end (* locale *) lemma transforms_tm_empty_or_endswithI [transforms_intros]: fixes j1 j2 :: tapeidx and z :: symbol fixes tps tps' :: "tape list" and k :: nat and zs :: "symbol list" assumes "j1 \ j2" "j1 < k" "j2 < k" and "length tps = k" and "proper_symbols zs" assumes "tps ! j1 = (\zs\, 1)" "tps ! j2 = (\0\\<^sub>N, 1)" assumes "ttt = 18 + 2 * length zs" assumes "tps' = tps [j2 := (\zs = [] \ last zs = z\\<^sub>B, 1)]" shows "transforms (tm_empty_or_endswith j1 j2 z) tps ttt tps'" proof - interpret loc: turing_machine_empty_or_endswith j1 j2 z . show ?thesis using assms loc.tps4_def loc.tm4 loc.tm4_eq_tm_empty_or_endswith by simp qed subsection \Stripping trailing symbols\ text \ Stripping the symbol $z$ from the end of a symbol sequence @{term zs} means: \ definition rstrip :: "symbol \ symbol list \ symbol list" where "rstrip z zs \ take (LEAST i. i \ length zs \ set (drop i zs) \ {z}) zs" lemma length_rstrip: "length (rstrip z zs) = (LEAST i. i \ length zs \ set (drop i zs) \ {z})" using rstrip_def wellorder_Least_lemma[where ?P="\i. i \ length zs \ set (drop i zs) \ {z}"] by simp lemma length_rstrip_le: "length (rstrip z zs) \ length zs" using rstrip_def by simp lemma rstrip_stripped: assumes "i \ length (rstrip z zs)" and "i < length zs" shows "zs ! i = z" proof - let ?P = "\i. i \ length zs \ set (drop i zs) \ {z}" have "?P (length zs)" by simp then have "?P i" using assms length_rstrip LeastI[where ?P="?P"] Least_le[where ?P="?P"] by (metis (mono_tags, lifting) dual_order.trans order_less_imp_le set_drop_subset_set_drop) then have "set (drop i zs) \ {z}" by simp then show ?thesis using assms(2) by (metis Cons_nth_drop_Suc drop_eq_Nil2 leD list.set(2) set_empty singleton_insert_inj_eq subset_singletonD) qed lemma rstrip_replicate: "rstrip z (replicate n z) = []" using rstrip_def by (metis (no_types, lifting) Least_eq_0 empty_replicate set_drop_subset set_replicate take_eq_Nil zero_le) lemma rstrip_not_ex: assumes "\ (\i z)" shows "rstrip z zs = []" using assms rstrip_def by (metis in_set_conv_nth replicate_eqI rstrip_replicate) lemma assumes "\i z" shows rstrip_ex_length: "length (rstrip z zs) > 0" and rstrip_ex_last: "last (rstrip z zs) \ z" proof - let ?P = "\i. i \ length zs \ set (drop i zs) \ {z}" obtain i where i: "i < length zs" "zs ! i \ z" using assms by auto then have "\ set (drop i zs) \ {z}" by (metis Cons_nth_drop_Suc drop_eq_Nil2 leD list.set(2) set_empty singleton_insert_inj_eq' subset_singletonD) then have "\ set (drop 0 zs) \ {z}" by (metis drop.simps(1) drop_0 set_drop_subset set_empty subset_singletonD) then show len: "length (rstrip z zs) > 0" using length_rstrip by (metis (no_types, lifting) LeastI bot.extremum drop_all dual_order.refl gr0I list.set(1)) let ?j = "length (rstrip z zs) - 1" have 3: "?j < length (rstrip z zs)" using len by simp then have 4: "?j < Least ?P" using length_rstrip by simp have 5: "?P (length (rstrip z zs))" using LeastI_ex[of "?P"] length_rstrip by fastforce show "last (rstrip z zs) \ z" proof (rule ccontr) assume "\ last (rstrip z zs) \ z" then have "last (rstrip z zs) = z" by simp then have "rstrip z zs ! ?j = z" using len by (simp add: last_conv_nth) then have 2: "zs ! ?j = z" using len length_rstrip rstrip_def by auto have "?P ?j" proof - have "?j \ length zs" using 3 length_rstrip_le by (meson le_eq_less_or_eq order_less_le_trans) moreover have "set (drop ?j zs) \ {z}" using 5 3 2 by (metis Cons_nth_drop_Suc One_nat_def Suc_pred insert_subset len list.simps(15) order_less_le_trans set_eq_subset) ultimately show ?thesis by simp qed then show False using 4 Least_le[of "?P"] by fastforce qed qed text \ A Turing machine stripping the non-blank, non-start symbol $z$ from a proper symbol sequence works in the obvious way. First it moves to the end of the symbol sequence, that is, to the first blank. Then it moves left to the first non-$z$ symbol thereby overwriting every symbol with a blank. Finally it performs a ``carriage return''. \ definition tm_rstrip :: "symbol \ tapeidx \ machine" where "tm_rstrip z j \ tm_right_until j {\} ;; tm_left j ;; tm_lconst_until j j (UNIV - {z}) \ ;; tm_cr j" lemma tm_rstrip_tm: assumes "k \ 2" and "G \ 4" and "0 < j" and "j < k" shows "turing_machine k G (tm_rstrip z j)" using assms tm_right_until_tm tm_left_tm tm_lconst_until_tm tm_cr_tm tm_rstrip_def by simp locale turing_machine_rstrip = fixes z :: symbol and j :: tapeidx begin definition "tm1 \ tm_right_until j {\}" definition "tm2 \ tm1 ;; tm_left j" definition "tm3 \ tm2 ;; tm_lconst_until j j (UNIV - {z}) \" definition "tm4 \ tm3 ;; tm_cr j" lemma tm4_eq_tm_rstrip: "tm4 = tm_rstrip z j" unfolding tm4_def tm3_def tm2_def tm1_def tm_rstrip_def by simp context fixes tps0 :: "tape list" and zs :: "symbol list" and k :: nat assumes z: "z > 1" assumes zs: "proper_symbols zs" assumes jk: "0 < j" "j < k" "length tps0 = k" assumes tps0: "tps0 ! j = (\zs\, 1)" begin definition "tps1 \ tps0 [j := (\zs\, Suc (length zs))]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (length zs)" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: tps0 tps1_def jk time: assms) have *: "tps0 ! j = (\zs\, 1)" using tps0 jk by simp show "rneigh (tps0 ! j) {\} (length zs)" using * zs by (intro rneighI) auto qed definition "tps2 \ tps0 [j := (\zs\, length zs)]" lemma tm2 [transforms_intros]: assumes "ttt = length zs + 2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: tps1_def tps2_def jk time: assms) definition "tps3 \ tps0 [j := (\rstrip z zs\, length (rstrip z zs))]" lemma tm3 [transforms_intros]: assumes "ttt = length zs + 2 + Suc (length zs - length (rstrip z zs))" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: tps2_def tps3_def jk time: assms jk tps2_def) let ?n = "length zs - length (rstrip z zs)" have *: "tps2 ! j = (\zs\, length zs)" using tps2_def jk by simp show "lneigh (tps2 ! j) (UNIV - {z}) ?n" proof (cases "\i z") case True then have 1: "length (rstrip z zs) > 0" using rstrip_ex_length by simp show ?thesis proof (rule lneighI) show "(tps2 ::: j) (tps2 :#: j - ?n) \ UNIV - {z}" using * 1 contents_inbounds True length_rstrip length_rstrip_le rstrip_def rstrip_ex_last by (smt (verit, best) DiffI One_nat_def UNIV_I diff_diff_cancel diff_less fst_conv last_conv_nth le_eq_less_or_eq length_greater_0_conv less_Suc_eq_le nth_take singletonD snd_conv) have "\n'. n' < ?n \ (tps2 ::: j) (tps2 :#: j - n') = z" using * rstrip_stripped by simp then show "\n'. n' < ?n \ (tps2 ::: j) (tps2 :#: j - n') \ UNIV - {z}" by simp qed next case False then have 1: "rstrip z zs = []" using rstrip_not_ex by simp show ?thesis proof (rule lneighI) show "(tps2 ::: j) (tps2 :#: j - ?n) \ UNIV - {z}" using * 1 z by simp show "\n'. n' < ?n \ (tps2 ::: j) (tps2 :#: j - n') \ UNIV - {z}" using * rstrip_stripped by simp qed qed have "lconstplant (\zs\, length zs) \ ?n = (\rstrip z zs\, length (rstrip z zs))" (is "?lhs = _") proof - have "?lhs = (\i. if length zs - ?n < i \ i \ length zs then \ else \zs\ i, length zs - ?n)" using lconstplant[of "(\zs\, length zs)" 0 "?n"] by auto moreover have "(\i. if length zs - ?n < i \ i \ length zs then \ else \zs\ i) = \rstrip z zs\" proof fix i consider "length zs - ?n < i \ i \ length zs" | "i > length zs" | "i \ length (rstrip z zs)" by linarith then show "(if length zs - ?n < i \ i \ length zs then \ else \zs\ i) = \rstrip z zs\ i" proof (cases) case 1 then show ?thesis by auto next case 2 then show ?thesis by (metis contents_outofbounds diff_diff_cancel length_rstrip_le less_imp_diff_less) next case 3 then show ?thesis using contents_def length_rstrip length_rstrip_le rstrip_def by auto qed qed moreover have "length zs - ?n = length (rstrip z zs)" using diff_diff_cancel length_rstrip_le by simp ultimately show ?thesis by simp qed then have "lconstplant (tps2 ! j) \ ?n = (\rstrip z zs\, length (rstrip z zs))" using tps2_def jk by simp then show "tps3 = tps2 [j := tps2 ! j |-| ?n, j := lconstplant (tps2 ! j) \ ?n]" unfolding tps3_def tps2_def by simp qed definition "tps4 \ tps0 [j := (\rstrip z zs\, 1)]" lemma tm4: assumes "ttt = length zs + 2 + Suc (length zs - length (rstrip z zs)) + length (rstrip z zs) + 2" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: tps3_def tps4_def jk time: assms tps3_def jk) show "clean_tape (tps3 ! j)" using tps3_def jk zs rstrip_def by simp qed lemma tm4': assumes "ttt = 3 * length zs + 5" shows "transforms tm4 tps0 ttt tps4" proof - let ?ttt = "length zs + 2 + Suc (length zs - length (rstrip z zs)) + length (rstrip z zs) + 2" have "?ttt = length zs + 5 + (length zs - length (rstrip z zs)) + length (rstrip z zs)" by simp also have "... \ length zs + 5 + length zs + length (rstrip z zs)" by simp also have "... \ length zs + 5 + length zs + length zs" using length_rstrip_le by simp also have "... = 3 * length zs + 5" by simp finally have "?ttt \ 3 * length zs + 5" . then show ?thesis using assms transforms_monotone tm4 by simp qed end (* context *) end (* locale *) lemma transforms_tm_rstripI [transforms_intros]: fixes z :: symbol and j :: tapeidx fixes tps tps' :: "tape list" and zs :: "symbol list" and k :: nat assumes "z > 1" and "0 < j" "j < k" and "proper_symbols zs" and "length tps = k" assumes "tps ! j = (\zs\, 1)" assumes "ttt = 3 * length zs + 5" assumes "tps' = tps[j := (\rstrip z zs\, 1)]" shows "transforms (tm_rstrip z j) tps ttt tps'" proof - interpret loc: turing_machine_rstrip z j . show ?thesis using assms loc.tm4' loc.tps4_def loc.tm4_eq_tm_rstrip by simp qed subsection \Writing arbitrary length sequences of the same symbol\ text \ The next Turing machine accepts a number $n$ on tape $j_1$ and writes the symbol sequence $z^n$ to tape $j_2$. The symbol $z$ is a parameter of the TM. The TM decrements the number on tape $j_1$ until it reaches zero. \ definition tm_write_replicate :: "symbol \ tapeidx \ tapeidx \ machine" where "tm_write_replicate z j1 j2 \ WHILE [] ; \rs. rs ! j1 \ \ DO tm_char j2 z ;; tm_decr j1 DONE ;; tm_cr j2" lemma tm_write_replicate_tm: assumes "0 < j1" and "0 < j2" and "j1 < k" and "j2 < k" and "j1 \ j2" and "G \ 4" and "z < G" shows "turing_machine k G (tm_write_replicate z j1 j2)" unfolding tm_write_replicate_def using turing_machine_loop_turing_machine Nil_tm tm_char_tm tm_decr_tm tm_cr_tm assms by simp locale turing_machine_write_replicate = fixes j1 j2 :: tapeidx and z :: symbol begin definition "tm1 \ tm_char j2 z" definition "tm2 \ tm1 ;; tm_decr j1" definition "tmL \ WHILE [] ; \rs. rs ! j1 \ \ DO tm2 DONE" definition "tm3 \ tmL ;; tm_cr j2" lemma tm3_eq_tm_write_replicate: "tm3 = tm_write_replicate z j1 j2" using tm3_def tm2_def tm1_def tm_write_replicate_def tmL_def by simp context fixes tps0 :: "tape list" and k n :: nat assumes jk: "length tps0 = k" "0 < j1" "0 < j2" "j1 \ j2" "j1 < k" "j2 < k" and z: "1 < z" assumes tps0: "tps0 ! j1 = (\n\\<^sub>N, 1)" "tps0 ! j2 = (\[]\, 1)" begin definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j1 := (\n - t\\<^sub>N, 1), j2 := (\replicate t z\, Suc t)]" lemma tpsL0: "tpsL 0 = tps0" using tpsL_def tps0 jk by (metis One_nat_def diff_zero list_update_id replicate_empty) definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0 [j1 := (\n - t\\<^sub>N, 1), j2 := (\replicate (Suc t) z\, Suc (Suc t))]" lemma tmL1 [transforms_intros]: "transforms tm1 (tpsL t) 1 (tpsL1 t)" unfolding tm1_def proof (tform tps: tpsL_def tpsL1_def tps0 jk) have "tpsL t :#: j2 = Suc t" using tpsL1_def jk by (metis length_list_update nth_list_update_eq snd_conv tpsL_def) moreover have "tpsL t ::: j2 = \replicate t z\" using tpsL1_def jk by (metis fst_conv length_list_update nth_list_update_eq tpsL_def) moreover have "\replicate t z\(Suc t := z) = \replicate (Suc t) z\" using contents_snoc by (metis length_replicate replicate_Suc replicate_append_same) ultimately show "tpsL1 t = (tpsL t)[j2 := tpsL t ! j2 |:=| z |+| 1]" unfolding tpsL1_def tpsL_def by simp qed lemma tmL2: assumes "ttt = 9 + 2 * nlength (n - t)" shows "transforms tm2 (tpsL t) ttt (tpsL (Suc t))" unfolding tm2_def proof (tform tps: assms tpsL_def tpsL1_def tps0 jk) show "tpsL (Suc t) = (tpsL1 t)[j1 := (\n - t - 1\\<^sub>N, 1)]" unfolding tpsL_def tpsL1_def using jk by (simp add: list_update_swap) qed lemma tmL2' [transforms_intros]: assumes "ttt = 9 + 2 * nlength n" shows "transforms tm2 (tpsL t) ttt (tpsL (Suc t))" proof - have "9 + 2 * nlength (n - t) \ 9 + 2 * nlength n" using nlength_mono[of "n - t" n] by simp then show ?thesis using assms tmL2 transforms_monotone by blast qed lemma tmL [transforms_intros]: assumes "ttt = n * (11 + 2 * nlength n) + 1" shows "transforms tmL (tpsL 0) ttt (tpsL n)" unfolding tmL_def proof (tform) let ?t = "9 + 2 * nlength n" show "\i. i < n \ read (tpsL i) ! j1 \ \" using jk tpsL_def read_ncontents_eq_0 by simp show "\ read (tpsL n) ! j1 \ \" using jk tpsL_def read_ncontents_eq_0 by simp show "n * (9 + 2 * nlength n + 2) + 1 \ ttt" using assms by simp qed definition tps3 :: "tape list" where "tps3 \ tps0 [j1 := (\0\\<^sub>N, 1), j2 := (\replicate n z\, 1)]" lemma tm3: assumes "ttt = n * (12 + 2 * nlength n) + 4" shows "transforms tm3 (tpsL 0) ttt tps3" unfolding tm3_def proof (tform tps: z tpsL_def tps3_def tps0 jk) have "ttt = Suc (n * (11 + 2 * nlength n)) + Suc (Suc (Suc n))" proof - have "Suc (n * (11 + 2 * nlength n)) + Suc (Suc (Suc n)) = n * (11 + 2 * nlength n) + 4 + n" by simp also have "... = n * (12 + 2 * nlength n) + 4" by algebra finally have "Suc (n * (11 + 2 * nlength n)) + Suc (Suc (Suc n)) = ttt" using assms by simp then show ?thesis by simp qed then show "ttt = n * (11 + 2 * nlength n) + 1 + (tpsL n :#: j2 + 2)" using tpsL_def jk by simp qed lemma tm3': assumes "ttt = n * (12 + 2 * nlength n) + 4" shows "transforms tm3 tps0 ttt tps3" using tm3 tpsL0 assms by simp end end lemma transforms_tm_write_replicateI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and ttt k n :: nat assumes "length tps = k" "0 < j1" "0 < j2" "j1 \ j2" "j1 < k" "j2 < k" and "1 < z" assumes "tps ! j1 = (\n\\<^sub>N, 1)" "tps ! j2 = (\[]\, 1)" assumes "ttt = n * (12 + 2 * nlength n) + 4" assumes "tps' = tps [j1 := (\0\\<^sub>N, 1), j2 := (\replicate n z\, 1)]" shows "transforms (tm_write_replicate z j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_write_replicate j1 j2 . show ?thesis using assms loc.tm3' loc.tps3_def loc.tm3_eq_tm_write_replicate by simp qed subsection \Extracting the elements of a pair\ text \ In Section~\ref{s:tm-basic-pair} we defined a pairing function for strings. For example, $\langle \bbbI\bbbI, \bbbO\bbbO\rangle$ is first mapped to $\bbbI\bbbI\#\bbbO\bbbO$ and ultimately represented as $\bbbO\bbbI\bbbO\bbbI\bbbI\bbbI\bbbO\bbbO\bbbO\bbbO$. A Turing machine that is to compute a function for the argument $\langle \bbbI\bbbI, \bbbO\bbbO\rangle$ would receive as input the symbols \textbf{0101110000}. Typically the TM would then extract the two components \textbf{11} and \textbf{00}. In this section we devise TMs to do just that. As it happens, applying the quaternary alphabet decoding function @{const bindecode} (see Section~\ref{s:tm-quaternary}) to such a symbol sequence gets us halfway to extracting the elements of the pair. For example, decoding \textbf{0101110000} yields @{text "\\\\\"}, and now the TM only has to locate the @{text \}. A Turing machine cannot rely on being given a well-formed pair. After decoding, the symbol sequence might have more or fewer than one @{text \} symbol or even @{text "\"} symbols. The following functions @{term first} and @{term second} are designed to extract the first and second element of a symbol sequence representing a pair, and for other symbol sequences at least allow for an efficient implementation. Implementations will come further down in this section. \ definition first :: "symbol list \ symbol list" where "first ys \ take (if \i {\, \} then LEAST i. i < length ys \ ys ! i \ {\, \} else length ys) ys" definition second :: "symbol list \ symbol list" where "second zs \ drop (Suc (length (first zs))) zs" lemma firstD: assumes "\i {\, \}" and "m = (LEAST i. i < length ys \ ys ! i \ {\, \})" shows "m < length ys" and "ys ! m \ {\, \}" and "\i {\, \}" using LeastI_ex[OF assms(1)] assms(2) by simp_all (use less_trans not_less_Least in blast) lemma firstI: assumes "m < length ys" and "ys ! m \ {\, \}" and "\i {\, \}" shows "(LEAST i. i < length ys \ ys ! i \ {\, \}) = m" using assms by (metis (mono_tags, lifting) LeastI less_linear not_less_Least) lemma length_first_ex: assumes "\i {\, \}" and "m = (LEAST i. i < length ys \ ys ! i \ {\, \})" shows "length (first ys) = m" proof - have "m < length ys" using assms firstD(1) by presburger moreover have "first ys = take m ys" using assms first_def by simp ultimately show ?thesis by simp qed lemma first_notex: assumes "\ (\i {\, \})" shows "first ys = ys" using assms first_def by auto lemma length_first: "length (first ys) \ length ys" using length_first_ex first_notex first_def by simp lemma length_second_first: "length (second zs) = length zs - Suc (length (first zs))" using second_def by simp lemma length_second: "length (second zs) \ length zs" using second_def by simp text \ Our next goal is to show that @{const first} and @{const second} really extract the first and second element of a pair. \ lemma bindecode_bitenc: fixes x :: string shows "bindecode (string_to_symbols (bitenc x)) = string_to_symbols x" proof (induction x) case Nil then show ?case using less_2_cases_iff by force next case (Cons a x) have "bitenc (a # x) = bitenc [a] @ bitenc x" by simp then have "string_to_symbols (bitenc (a # x)) = string_to_symbols (bitenc [a] @ bitenc x)" by simp then have "string_to_symbols (bitenc (a # x)) = string_to_symbols (bitenc [a]) @ string_to_symbols (bitenc x)" by simp then have "bindecode (string_to_symbols (bitenc (a # x))) = bindecode (string_to_symbols (bitenc [a]) @ string_to_symbols (bitenc x))" by simp also have "... = bindecode (string_to_symbols (bitenc [a])) @ bindecode (string_to_symbols (bitenc x))" using bindecode_append length_bitenc by (metis (no_types, lifting) dvd_triv_left length_map) also have "... = bindecode (string_to_symbols (bitenc [a])) @ string_to_symbols x" using Cons by simp also have "... = string_to_symbols [a] @ string_to_symbols x" using bindecode_def by simp also have "... = string_to_symbols ([a] @ x)" by simp also have "... = string_to_symbols (a # x)" by simp finally show ?case . qed lemma bindecode_string_pair: fixes x u :: string shows "bindecode \x; u\ = string_to_symbols x @ [\] @ string_to_symbols u" proof - have "bindecode \x; u\ = bindecode (string_to_symbols (bitenc x @ [True, True] @ bitenc u))" using string_pair_def by simp also have "... = bindecode (string_to_symbols (bitenc x) @ string_to_symbols [\, \] @ string_to_symbols (bitenc u))" by simp also have "... = bindecode (string_to_symbols (bitenc x)) @ bindecode (string_to_symbols [\, \]) @ bindecode (string_to_symbols (bitenc u))" proof - have "even (length (string_to_symbols [True, True]))" by simp moreover have "even (length (string_to_symbols (bitenc y)))" for y by (simp add: length_bitenc) ultimately show ?thesis using bindecode_append length_bindecode length_bitenc - by (smt (z3) add_mult_distrib2 add_self_div_2 dvd_triv_left length_append length_map mult_2) + by (smt (verit) add_mult_distrib2 add_self_div_2 dvd_triv_left length_append length_map mult_2) qed also have "... = string_to_symbols x @ bindecode (string_to_symbols [\, \]) @ string_to_symbols u" using bindecode_bitenc by simp also have "... = string_to_symbols x @ [\] @ string_to_symbols u" using bindecode_def by simp finally show ?thesis . qed lemma first_pair: fixes ys :: "symbol list" and x u :: string assumes "ys = bindecode \x; u\" shows "first ys = string_to_symbols x" proof - have ys: "ys = string_to_symbols x @ [\] @ string_to_symbols u" using bindecode_string_pair assms by simp have bs: "bit_symbols (string_to_symbols x)" by simp have "ys ! (length (string_to_symbols x)) = \" using ys by (metis append_Cons nth_append_length) then have ex: "ys ! (length (string_to_symbols x)) \ {\, \}" by simp have "(LEAST i. i < length ys \ ys ! i \ {\, \}) = length (string_to_symbols x)" using ex ys bs by (intro firstI) (simp_all add: nth_append) moreover have "length (string_to_symbols x) < length ys" using ys by simp ultimately have "first ys = take (length (string_to_symbols x)) ys" using ex first_def by auto then show "first ys = string_to_symbols x" using ys by simp qed lemma second_pair: fixes ys :: "symbol list" and x u :: string assumes "ys = bindecode \x; u\" shows "second ys = string_to_symbols u" proof - have ys: "ys = string_to_symbols x @ [\] @ string_to_symbols u" using bindecode_string_pair assms by simp let ?m = "length (string_to_symbols x)" have "length (first ys) = ?m" using assms first_pair by presburger moreover have "drop (Suc ?m) ys = string_to_symbols u" using ys by simp ultimately have "drop (Suc (length (first ys))) ys = string_to_symbols u" by simp then show ?thesis using second_def by simp qed subsubsection \A Turing machine for extracting the first element\ text \ Unlike most other Turing machines, the one in this section is not meant to be reusable, but rather to compute a function, namely the function @{const first}. For this reason there are no tape index parameters. Instead, the encoded pair is expected on the input tape, and the output is written to the output tape. \null \ lemma bit_symbols_first: assumes "ys = bindecode (string_to_symbols x)" shows "bit_symbols (first ys)" proof (cases "\i {\, \}") case True define m where "m = (LEAST i. i < length ys \ ys ! i \ {\, \})" then have m: "m < length ys" "ys ! m \ {\, \}" "\i {\, \}" using firstD[OF True] by blast+ have len: "length (first ys) = m" using length_first_ex[OF True] m_def by simp have "bit_symbols (string_to_symbols x)" by simp then have "\i {2..<6}" using assms bindecode2345 by simp then have "\i {2..<6}" using m(1) by simp then have "\i {2..<4}" using m(3) by fastforce then show ?thesis using first_def len by auto next case False then have 1: "\i {\, \}" by simp have "bit_symbols (string_to_symbols x)" by simp then have "\i {2..<6}" using assms bindecode2345 by simp then have "\i {2..<4}" using 1 by fastforce then show ?thesis using False first_notex by auto qed definition tm_first :: machine where "tm_first \ tm_right_many {0, 1, 2} ;; tm_bindecode 0 2 ;; tm_cp_until 2 1 {\, \, \}" lemma tm_first_tm: "G \ 6 \ k \ 3 \ turing_machine k G tm_first" unfolding tm_first_def using tm_cp_until_tm tm_start_tm tm_bindecode_tm tm_right_many_tm by simp locale turing_machine_fst_pair = fixes k :: nat assumes k: "k \ 3" begin definition "tm1 \ tm_right_many {0, 1, 2}" definition "tm2 \ tm1 ;; tm_bindecode 0 2" definition "tm3 \ tm2 ;; tm_cp_until 2 1 {\, \, \}" lemma tm3_eq_tm_first: "tm3 = tm_first" using tm1_def tm2_def tm3_def tm_first_def by simp context fixes xs :: "symbol list" assumes bs: "bit_symbols xs" begin definition "tps0 \ snd (start_config k xs)" lemma lentps [simp]: "length tps0 = k" using tps0_def start_config_length k by simp lemma tps0_0: "tps0 ! 0 = (\xs\, 0)" using tps0_def start_config_def contents_def by auto lemma tps0_gt_0: "j > 0 \ j < k \ tps0 ! j = (\[]\, 0)" using tps0_def start_config_def contents_def by auto definition "tps1 \ tps0 [0 := (\xs\, 1), 1 := (\[]\, 1), 2 := (\[]\, 1)]" lemma tm1 [transforms_intros]: "transforms tm1 tps0 1 tps1" unfolding tm1_def proof (tform) show "tps1 = map (\j. if j \ {0, 1, 2} then tps0 ! j |+| 1 else tps0 ! j) [0.. tps0 [0 := (\xs\, 1), 1 := (\[]\, 1), 2 := (\bindecode xs\, 1)]" lemma tm2 [transforms_intros]: assumes "ttt = 8 + 3 * length xs" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def by (tform tps: bs k tps1_def assms tps2_def) definition "tps3 \ tps0 [0 := (\xs\, 1), 1 := (\first (bindecode xs)\, Suc (length (first (bindecode xs)))), 2 := (\bindecode xs\, Suc (length (first (bindecode xs))))]" lemma tm3: assumes "ttt = 8 + 3 * length xs + Suc (length (first (bindecode xs)))" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: k tps2_def time: assms) let ?ys = "bindecode xs" have tps2: "tps2 ! 2 = (\?ys\, 1)" using tps2_def k by simp show "rneigh (tps2 ! 2) {\, \, \} (length (first ?ys))" proof (cases "\i {\, \}") case ex5: True define m where "m = (LEAST i. i < length ?ys \ ?ys ! i \ {\, \})" then have m: "m = length (first ?ys)" using length_first_ex ex5 by simp show ?thesis proof (rule rneighI) have "?ys ! m \ {\, \}" using firstD m_def ex5 by blast then show "(tps2 ::: 2) (tps2 :#: 2 + length (first ?ys)) \ {\, \, \}" using m tps2 contents_def by simp show "(tps2 ::: 2) (tps2 :#: 2 + i) \ {\, \, \}" if "i < length (first ?ys)" for i proof - have "m < length ?ys" using ex5 firstD(1) length_first_ex m by blast then have "length (first ?ys) < length ?ys" using m by simp then have "i < length ?ys" using that by simp then have "?ys ! i \ 0" using proper_bindecode by fastforce moreover have "?ys ! i \ {\, \}" using ex5 firstD(3) length_first_ex that by blast ultimately show ?thesis using Suc_neq_Zero \i < length (bindecode xs)\ tps2 by simp qed qed next case notex5: False then have ys: "?ys = first ?ys" using first_notex by simp show ?thesis proof (rule rneighI) show "(tps2 ::: 2) (tps2 :#: 2 + length (first ?ys)) \ {\, \, \}" using ys tps2 by simp show "(tps2 ::: 2) (tps2 :#: 2 + i) \ {\, \, \}" if "i < length (first ?ys)" for i using notex5 that ys proper_bindecode contents_inbounds by (metis Suc_leI add_gr_0 diff_Suc_1 fst_conv gr_implies_not0 insert_iff plus_1_eq_Suc snd_conv tps2 zero_less_one) qed qed show "tps3 = tps2[2 := tps2 ! 2 |+| length (first ?ys), 1 := implant (tps2 ! 2) (tps2 ! 1) (length (first ?ys))]" (is "_ = ?tps") proof - have 0: "tps3 ! 0 = ?tps ! 0" using tps2_def tps3_def by simp have 1: "tps3 ! 2 = ?tps ! 2" using tps2_def tps3_def k by simp have lentps2: "length tps2 > 2" using k tps2_def by simp have "implant (tps2 ! 2) (tps2 ! 1) (length (first ?ys)) = (\first ?ys\, Suc (length (first ?ys)))" proof - have len: "length (first ?ys) \ length ?ys" using first_def by simp have "tps2 ! 1 = (\[]\, 1)" using tps2_def lentps2 by simp then have "implant (tps2 ! 2) (tps2 ! 1) (length (first ?ys)) = implant (\?ys\, 1) (\[]\, 1) (length (first ?ys))" using tps2 by simp also have "... = (\take (length (first ?ys)) ?ys\, Suc (length (first ?ys)))" using implant_contents[of 1 "length (first ?ys)" ?ys "[]"] len by simp also have "... = (\first ?ys\, Suc (length (first ?ys)))" using first_def using first_notex length_first_ex by presburger finally show ?thesis . qed moreover have "length tps2 > 2" using k tps2_def by simp ultimately show ?thesis using 0 1 tps2_def tps3_def tps0_def lentps k tps2 - by (smt (z3) length_list_update list_update_overwrite list_update_swap nth_list_update) + by (smt (verit) length_list_update list_update_overwrite list_update_swap nth_list_update) qed qed lemma tm3': assumes "ttt = 9 + 4 * length xs" shows "transforms tm3 tps0 ttt tps3" proof - let ?t = "8 + 3 * length xs + Suc (length (first (bindecode xs)))" have "?t \ 8 + 3 * length xs + Suc (length (bindecode xs))" using length_first by (meson Suc_le_mono add_le_mono order_refl) also have "... \ 8 + 3 * length xs + Suc (length xs)" using length_bindecode by simp also have "... = 9 + 3 * length xs + length xs" by simp also have "... = 9 + 4 * length xs" by simp finally have "?t \ ttt" using assms(1) by simp moreover have "transforms tm3 tps0 ?t tps3" using tm3 by simp ultimately show ?thesis using transforms_monotone by simp qed end (* context tps *) lemma tm3_computes: "computes_in_time k tm3 (\x. symbols_to_string (first (bindecode (string_to_symbols x)))) (\n. 9 + 4 * n)" proof - define f where "f = (\x. symbols_to_string (first (bindecode (string_to_symbols x))))" define T :: "nat \ nat" where "T = (\n. 9 + 4 * n)" have "computes_in_time k tm3 f T" proof fix x :: string let ?xs = "string_to_symbols x" have bs: "bit_symbols ?xs" by simp define tps where "tps = tps3 ?xs" have trans: "transforms tm3 (tps0 ?xs) (9 + 4 * length ?xs) tps" using bs tm3' tps_def by blast have "tps3 ?xs ::: 1 = \first (bindecode ?xs)\" using bs tps3_def k by simp moreover have "bit_symbols (first (bindecode ?xs))" using bit_symbols_first by simp ultimately have "tps3 ?xs ::: 1 = string_to_contents (symbols_to_string (first (bindecode ?xs)))" using bit_symbols_to_symbols contents_string_to_contents by simp then have *: "tps ::: 1 = string_to_contents (f x)" using tps_def f_def by auto then have "transforms tm3 (snd (start_config k (string_to_symbols x))) (T (length x)) tps" using trans T_def tps0_def by simp then show "\tps. tps ::: 1 = string_to_contents (f x) \ transforms tm3 (snd (start_config k (string_to_symbols x))) (T (length x)) tps" using * by auto qed then show ?thesis using f_def T_def by simp qed end (* locale turing_machine_fst_pair *) lemma tm_first_computes: assumes "k \ 3" shows "computes_in_time k tm_first (\x. symbols_to_string (first (bindecode (string_to_symbols x)))) (\n. 9 + 4 * n)" proof - interpret loc: turing_machine_fst_pair k using turing_machine_fst_pair.intro assms by simp show ?thesis using loc.tm3_eq_tm_first loc.tm3_computes by simp qed subsubsection \A Turing machine for splitting pairs\ text \ The next Turing machine expects a proper symbol sequence @{term zs} on tape $j_1$ and outputs @{term "first zs"} and @{term "second zs"} on tapes $j_2$ and $j_3$, respectively. \ definition tm_unpair :: "tapeidx \ tapeidx \ tapeidx \ machine" where "tm_unpair j1 j2 j3 \ tm_cp_until j1 j2 {\, \, \} ;; tm_right j1 ;; tm_cp_until j1 j3 {\} ;; tm_cr j1 ;; tm_cr j2 ;; tm_cr j3" lemma tm_unpair_tm: assumes "k \ 2" and "G \ 4" and "0 < j2" and "0 < j3" and "j1 < k" "j2 < k" "j3 < k" shows "turing_machine k G (tm_unpair j1 j2 j3)" using tm_cp_until_tm tm_right_tm tm_cr_tm assms tm_unpair_def by simp locale turing_machine_unpair = fixes j1 j2 j3 :: tapeidx begin definition "tm1 \ tm_cp_until j1 j2 {\, \, \}" definition "tm2 \ tm1 ;; tm_right j1" definition "tm3 \ tm2 ;; tm_cp_until j1 j3 {\}" definition "tm4 \ tm3 ;; tm_cr j1" definition "tm5 \ tm4 ;; tm_cr j2" definition "tm6 \ tm5 ;; tm_cr j3" lemma tm6_eq_tm_unpair: "tm6 = tm_unpair j1 j2 j3" unfolding tm6_def tm5_def tm4_def tm3_def tm2_def tm1_def tm_unpair_def by simp context fixes tps0 :: "tape list" and k :: nat and zs :: "symbol list" assumes jk: "0 < j2" "0 < j3" "j1 \ j2" "j1 \ j3" "j2 \ j3" "j1 < k" "j2 < k" "j3 < k" "length tps0 = k" and zs: "proper_symbols zs" and tps0: "tps0 ! j1 = (\zs\, 1)" "tps0 ! j2 = (\[]\, 1)" "tps0 ! j3 = (\[]\, 1)" begin definition "tps1 \ tps0 [j1 := (\zs\, Suc (length (first zs))), j2 := (\first zs\, Suc (length (first zs)))]" lemma tm1 [transforms_intros]: assumes "ttt = Suc (length (first zs))" shows "transforms tm1 tps0 ttt tps1" unfolding tm1_def proof (tform tps: assms tps0 tps1_def jk) let ?n = "length (first zs)" have *: "tps0 ! j1 = (\zs\, 1)" using tps0 jk by simp show "rneigh (tps0 ! j1) {\, \, \} (length (first zs))" proof (cases "\i {\, \}") case ex5: True define m where "m = (LEAST i. i < length zs \ zs ! i \ {\, \})" then have m: "m = length (first zs)" using length_first_ex ex5 by simp show ?thesis proof (rule rneighI) have "zs ! m \ {\, \}" using firstD m_def ex5 by blast then show "(tps0 ::: j1) (tps0 :#: j1 + length (first zs)) \ {\, \, \}" using m * contents_def by simp show "(tps0 ::: j1) (tps0 :#: j1 + i) \ {\, \, \}" if "i < length (first zs)" for i proof - have "m < length zs" using ex5 firstD(1) length_first_ex m by blast then have "length (first zs) < length zs" using m by simp then have "i < length zs" using that by simp then have "zs ! i \ \" using zs by fastforce moreover have "zs ! i \ {\, \}" using ex5 firstD(3) length_first_ex that by blast ultimately show ?thesis using Suc_neq_Zero `i < length zs` * by simp qed qed next case notex5: False then have ys: "zs = first zs" using first_notex by simp show ?thesis proof (rule rneighI) show "(tps0 ::: j1) (tps0 :#: j1 + length (first zs)) \ {\, \, \}" using ys * by simp show "(tps0 ::: j1) (tps0 :#: j1 + i) \ {\, \, \}" if "i < length (first zs)" for i using notex5 that ys proper_bindecode contents_inbounds * zs by auto qed qed have 1: "implant (tps0 ! j1) (tps0 ! j2) ?n = (\first zs\, Suc ?n)" proof - have "implant (tps0 ! j1) (tps0 ! j2) ?n = (\[] @ take (length (first zs)) (drop (1 - 1) zs)\, Suc (length []) + length (first zs))" using implant_contents[of 1 "length (first zs)" zs "[]"] tps0(1,2) by (metis (mono_tags, lifting) add.right_neutral diff_Suc_1 le_eq_less_or_eq firstD(1) first_notex length_first_ex less_one list.size(3) plus_1_eq_Suc) then have "implant (tps0 ! j1) (tps0 ! j2) ?n = (\take ?n zs\, Suc ?n)" by simp then show "implant (tps0 ! j1) (tps0 ! j2) ?n = (\first zs\, Suc ?n)" using first_def length_first_ex by auto qed have 2: "tps0 ! j1 |+| ?n = (\zs\, Suc ?n)" using tps0 jk by simp show "tps1 = tps0 [j1 := tps0 ! j1 |+| ?n, j2 := implant (tps0 ! j1) (tps0 ! j2) ?n]" unfolding tps1_def using jk 1 2 by simp qed definition "tps2 \ tps0 [j1 := (\zs\, length (first zs) + 2), j2 := (\first zs\, Suc (length (first zs)))]" lemma tm2 [transforms_intros]: assumes "ttt = length (first zs) + 2" shows "transforms tm2 tps0 ttt tps2" unfolding tm2_def proof (tform tps: tps1_def jk tps2_def time: assms) have "tps1 ! j1 |+| 1 = (\zs\, length (first zs) + 2)" using tps1_def jk by simp then show "tps2 = tps1[j1 := tps1 ! j1 |+| 1]" unfolding tps2_def tps1_def using jk by (simp add: list_update_swap) qed definition "tps3 \ tps0 [j1 := (\zs\, length (first zs) + 2 + (length zs - Suc (length (first zs)))), j2 := (\first zs\, Suc (length (first zs))), j3 := (\second zs\, Suc (length (second zs)))]" lemma tm3 [transforms_intros]: assumes "ttt = length (first zs) + 2 + Suc (length zs - Suc (length (first zs)))" shows "transforms tm3 tps0 ttt tps3" unfolding tm3_def proof (tform tps: assms tps2_def tps3_def jk) let ?ll = "length (first zs)" let ?n = "length zs - Suc ?ll" have at_j1: "tps2 ! j1 = (\zs\, length (first zs) + 2)" using tps2_def jk by simp show "rneigh (tps2 ! j1) {0} ?n" proof (rule rneighI) show "(tps2 ::: j1) (tps2 :#: j1 + (length zs - Suc ?ll)) \ {0}" using at_j1 by simp show "(tps2 ::: j1) (tps2 :#: j1 + m) \ {0}" if "m < length zs - Suc ?ll" for m proof - have *: "(tps2 ::: j1) (tps2 :#: j1 + m) = \zs\ (?ll + 2 + m)" using at_j1 by simp have "Suc ?ll < length zs" using that by simp then have "?ll + 2 + m \ Suc (length zs)" using that by simp then have "\zs\ (?ll + 2 + m) = zs ! (?ll + 1 + m)" using that by simp then have "\zs\ (?ll + 2 + m) > 0" using zs that by (metis add.commute gr0I less_diff_conv not_add_less2 plus_1_eq_Suc) then show ?thesis using * by simp qed qed have 1: "implant (tps2 ! j1) (tps2 ! j3) ?n = (\second zs\, Suc (length (second zs)))" proof (cases "Suc ?ll \ length zs") case True have "implant (tps2 ! j1) (tps2 ! j3) ?n = implant (\zs\, ?ll + 2) (\[]\, 1) ?n" using tps2_def jk by (metis at_j1 nth_list_update_neq' tps0(3)) also have "... = (\take ?n (drop (Suc ?ll) zs)\, Suc ?n)" using True implant_contents by (metis (no_types, lifting) One_nat_def add.commute add_2_eq_Suc' append.simps(1) diff_Suc_1 dual_order.refl le_add_diff_inverse2 list.size(3) plus_1_eq_Suc zero_less_Suc) also have "... = (\take (length (second zs)) (drop (Suc ?ll) zs)\, Suc (length (second zs)))" using length_second_first by simp also have "... = (\second zs\, Suc (length (second zs)))" using second_def by simp finally show ?thesis . next case False then have "?n = 0" by simp then have "implant (tps2 ! j1) (tps2 ! j3) ?n = implant (\zs\, ?ll + 2) (\[]\, 1) 0" using tps2_def jk by (metis at_j1 nth_list_update_neq' tps0(3)) then have "implant (tps2 ! j1) (tps2 ! j3) ?n = (\[]\, 1)" using transplant_0 by simp moreover have "second zs = []" using False second_def by simp ultimately show ?thesis by simp qed show "tps3 = tps2 [j1 := tps2 ! j1 |+| ?n, j3 := implant (tps2 ! j1) (tps2 ! j3) ?n]" using tps3_def tps2_def using 1 jk at_j1 by (simp add: list_update_swap[of j1]) qed definition "tps4 \ tps0 [j1 := (\zs\, 1), j2 := (\first zs\, Suc (length (first zs))), j3 := (\second zs\, Suc (length (second zs)))]" lemma tm4: assumes "ttt = 2 * length (first zs) + 7 + 2 * (length zs - Suc (length (first zs)))" shows "transforms tm4 tps0 ttt tps4" unfolding tm4_def proof (tform tps: assms tps3_def tps4_def jk zs) have "tps3 ! j1 |#=| 1 = (\zs\, 1)" using tps3_def jk by simp then show "tps4 = tps3[j1 := tps3 ! j1 |#=| 1]" unfolding tps4_def tps3_def using jk by (simp add: list_update_swap) qed lemma tm4' [transforms_intros]: assumes "ttt = 4 * length zs + 7" shows "transforms tm4 tps0 ttt tps4" proof - have "2 * length (first zs) + 7 + 2 * (length zs - Suc (length (first zs))) \ 2 * length (first zs) + 7 + 2 * length zs" by simp also have "... \ 2 * length zs + 7 + 2 * length zs" using length_first by simp also have "... = ttt" using assms by simp finally have "2 * length (first zs) + 7 + 2 * (length zs - Suc (length (first zs))) \ ttt" . then show ?thesis using assms tm4 transforms_monotone by simp qed definition "tps5 \ tps0 [j1 := (\zs\, 1), j2 := (\first zs\, 1), j3 := (\second zs\, Suc (length (second zs)))]" lemma tm5 [transforms_intros]: assumes "ttt = 4 * length zs + 9 + Suc (length (first zs))" shows "transforms tm5 tps0 ttt tps5" unfolding tm5_def proof (tform tps: assms tps4_def tps5_def jk) show "clean_tape (tps4 ! j2)" using zs first_def tps4_def jk by simp have "tps4 ! j2 |#=| 1 = (\first zs\, 1)" using tps4_def jk by simp then show "tps5 = tps4[j2 := tps4 ! j2 |#=| 1]" unfolding tps5_def tps4_def using jk by (simp add: list_update_swap) qed definition "tps6 \ tps0 [j1 := (\zs\, 1), j2 := (\first zs\, 1), j3 := (\second zs\, 1)]" lemma tm6: assumes "ttt = 4 * length zs + 11 + Suc (length (first zs)) + Suc (length (second zs))" shows "transforms tm6 tps0 ttt tps6" unfolding tm6_def proof (tform tps: assms tps5_def tps6_def jk) show "clean_tape (tps5 ! j3)" using zs second_def tps5_def jk by simp qed definition "tps6' \ tps0 [j2 := (\first zs\, 1), j3 := (\second zs\, 1)]" lemma tps6': "tps6' = tps6" using tps6_def tps6'_def list_update_id tps0(1) by metis lemma tm6': assumes "ttt = 6 * length zs + 13" shows "transforms tm6 tps0 ttt tps6'" proof - have "4 * length zs + 11 + Suc (length (first zs)) + Suc (length (second zs)) \ 4 * length zs + 13 + length zs + length (second zs)" using length_first by simp also have "... \ 6 * length zs + 13" using length_second by simp finally have "4 * length zs + 11 + Suc (length (first zs)) + Suc (length (second zs)) \ ttt" using assms by simp then show ?thesis using tm6 tps6' transforms_monotone by simp qed end (* context *) end (* locale *) lemma transforms_tm_unpairI [transforms_intros]: fixes j1 j2 j3 :: tapeidx fixes tps tps' :: "tape list" and k :: nat and zs :: "symbol list" assumes "0 < j2" "0 < j3" "j1 \ j2" "j1 \ j3" "j2 \ j3" "j1 < k" "j2 < k" "j3 < k" and "length tps = k" and "proper_symbols zs" assumes "tps ! j1 = (\zs\, 1)" "tps ! j2 = (\[]\, 1)" "tps ! j3 = (\[]\, 1)" assumes "ttt = 6 * length zs + 13" assumes "tps' = tps [j2 := (\first zs\, 1), j3 := (\second zs\, 1)]" shows "transforms (tm_unpair j1 j2 j3) tps ttt tps'" proof - interpret loc: turing_machine_unpair j1 j2 j3 . show ?thesis using assms loc.tps6'_def loc.tm6' loc.tm6_eq_tm_unpair by metis qed end diff --git a/thys/Cook_Levin/Two_Four_Symbols.thy b/thys/Cook_Levin/Two_Four_Symbols.thy --- a/thys/Cook_Levin/Two_Four_Symbols.thy +++ b/thys/Cook_Levin/Two_Four_Symbols.thy @@ -1,1514 +1,1515 @@ section \Mapping between a binary and a quaternary alphabet\label{s:tm-quaternary}\ theory Two_Four_Symbols imports Arithmetic begin text \ Functions are defined over bit strings. For Turing machines these bits are represented by the symbols \textbf{0} and \textbf{1}. Sometimes we want a TM to receive a pair of bit strings or output a list of numbers. Or we might want the TM to interpret the input as a list of lists of numbers. All these objects can naturally be represented over a four-symbol (quaternary) alphabet, as we have seen for pairs in Section~\ref{s:tm-basic-pair} and for the lists in Sections~\ref{s:tm-numlist} and~\ref{s:tm-numlistlist}. To accommodate the aforementioned use cases, we define a straightforward mapping between the binary alphabet \textbf{01} and the quaternary alphabet @{text "\\\\"} and devise Turing machines to encode and decode symbol sequences. \ subsection \Encoding and decoding\label{s:tm-quaternary-encoding}\ text \ The encoding maps: \begin{tabular}{lcl} @{text \} & $\mapsto$ & @{text "\\"}\\ @{text \} & $\mapsto$ & @{text "\\"}\\ @{text "\"} & $\mapsto$ & @{text "\\"}\\ @{text \} & $\mapsto$ & @{text "\\"}\\ \end{tabular} \ text \ For example, the list $[6,0,1]$ is represented by the symbols @{text "\\\\\\\"}, which is encoded as @{text "\\\\\\\\\\\\\\"}. Pairing this sequence with the symbol sequence @{text "\\\\"} yields @{text "\\\\\\\\\\\\\\\\\\\"}, which is encoded as @{text "\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\\"}. \null \ definition binencode :: "symbol list \ symbol list" where "binencode ys \ concat (map (\y. [tosym ((y - 2) div 2), tosym ((y - 2) mod 2)]) ys)" lemma length_binencode [simp]: "length (binencode ys) = 2 * length ys" using binencode_def by (induction ys) simp_all lemma binencode_snoc: "binencode (zs @ [\]) = binencode zs @ [\, \]" "binencode (zs @ [\]) = binencode zs @ [\, \]" "binencode (zs @ [\]) = binencode zs @ [\, \]" "binencode (zs @ [\]) = binencode zs @ [\, \]" using binencode_def by simp_all lemma binencode_at_even: assumes "i < length ys" shows "binencode ys ! (2 * i) = 2 + (ys ! i - 2) div 2" using assms proof (induction ys arbitrary: i) case Nil then show ?case by simp next case (Cons y ys) have *: "binencode (y # ys) = [2 + (y - 2) div 2, 2 + (y - 2) mod 2] @ binencode ys" using binencode_def by simp show ?case proof (cases "i = 0") case True then show ?thesis using * by simp next case False then have "binencode (y # ys) ! (2 * i) = binencode ys ! (2 * i - 2)" using * by (metis One_nat_def length_Cons less_one list.size(3) mult_2 nat_mult_less_cancel_disj nth_append numerals(2) plus_1_eq_Suc) also have "... = binencode ys ! (2 * (i - 1))" using False by (simp add: right_diff_distrib') also have "... = 2 + (ys ! (i - 1) - 2) div 2" using False Cons by simp also have "... = 2 + ((y # ys) ! i - 2) div 2" using False by simp finally show ?thesis . qed qed lemma binencode_at_odd: assumes "i < length ys" shows "binencode ys ! (2 * i + 1) = 2 + (ys ! i - 2) mod 2" using assms proof (induction ys arbitrary: i) case Nil then show ?case by simp next case (Cons y ys) have *: "binencode (y # ys) = [2 + (y - 2) div 2, 2 + (y - 2) mod 2] @ binencode ys" using binencode_def by simp show ?case proof (cases "i = 0") case True then show ?thesis using * by simp next case False then have "binencode (y # ys) ! (2 * i + 1) = binencode ys ! (2 * i + 1 - 2)" using * by simp also have "... = binencode ys ! (2 * (i - 1) + 1)" using False by (metis Nat.add_diff_assoc2 One_nat_def Suc_leI diff_mult_distrib2 mult_2 mult_le_mono2 nat_1_add_1 neq0_conv) also have "... = 2 + (ys ! (i - 1) - 2) mod 2" using False Cons by simp also have "... = 2 + ((y # ys) ! i - 2) mod 2" using False by simp finally show ?thesis . qed qed text \ While @{const binencode} is defined for arbitrary symbol sequences, we only consider sequences over @{text "\\\\"} to be binencodable. \ abbreviation binencodable :: "symbol list \ bool" where "binencodable ys \ \i ys ! i \ ys ! i < 6" lemma binencodable_append: assumes "binencodable xs" and "binencodable ys" shows "binencodable (xs @ ys)" using assms prop_list_append by (simp add: nth_append) lemma bit_symbols_binencode: assumes "binencodable ys" shows "bit_symbols (binencode ys)" proof - have "2 \ binencode ys ! i \ binencode ys ! i \ 3" if "i < length (binencode ys)" for i proof (cases "even i") case True then have len: "i div 2 < length ys" using length_binencode that by simp moreover have "2 * (i div 2) = i" using True by simp ultimately have "binencode ys ! i = 2 + (ys ! (i div 2) - 2) div 2" using binencode_at_even[of "i div 2" ys] by simp moreover have "2 \ ys ! (i div 2) \ ys ! (i div 2) < 6" using len assms by simp ultimately show ?thesis by auto next case False then have len: "i div 2 < length ys" using length_binencode that by simp moreover have "2 * (i div 2) + 1 = i" using False by simp ultimately have "binencode ys ! i = 2 + (ys ! (i div 2) - 2) mod 2" using binencode_at_odd[of "i div 2" ys] by simp moreover have "2 \ ys ! (i div 2) \ ys ! (i div 2) < 6" using len assms by simp ultimately show ?thesis by simp qed then show ?thesis by fastforce qed text \ An encoded symbol sequence is of even length. When decoding a symbol sequence of odd length, we ignore the last symbol. For example, @{text "\\\\\\"} and @{text "\\\\\\\"} are both decoded to @{text "\\\"}. The bit symbol sequence @{term "[a, b]"} is decoded to this symbol: \ abbreviation decsym :: "symbol \ symbol \ symbol" where "decsym a b \ tosym (2 * todigit a + todigit b)" definition bindecode :: "symbol list \ symbol list" where "bindecode zs \ map (\i. decsym (zs ! (2 * i)) (zs ! (Suc (2 * i)))) [0..i {2..<6}" using assms bindecode_at by simp lemma bindecode_odd: assumes "length zs = 2 * n + 1" shows "bindecode zs = bindecode (take (2 * n) zs)" proof - have 1: "take (2 * n) zs ! (2 * i) = zs ! (2 * i)" if "i < n" for i using assms that by simp have 2: "take (2 * n) zs ! (Suc (2 * i)) = zs ! (Suc (2 * i))" if "i < n" for i using assms that by simp have "bindecode (take (2 * n) zs) = map (\i. decsym ((take (2 * n) zs) ! (2 * i)) ((take (2 * n) zs) ! (Suc (2 * i)))) [0..i. decsym ((take (2 * n) zs) ! (2 * i)) ((take (2 * n) zs) ! (Suc (2 * i)))) [0..i. decsym (zs ! (2 * i)) (zs ! (Suc (2 * i)))) [0..i. decsym (zs ! (2 * i)) (zs ! (Suc (2 * i)))) [0.. ys ! i = 3 \ ys ! i = 4 \ ys ! i = 5" using assms len - by (smt (z3) Suc_1 add_Suc_shift add_cancel_right_left eval_nat_numeral(3) + by (smt (verit) Suc_1 add_Suc_shift add_cancel_right_left eval_nat_numeral(3) less_Suc_eq numeral_3_eq_3 numeral_Bit0 verit_comp_simplify1(3)) then show ?thesis by auto qed finally show "bindecode ?zs ! i = ys ! i" . qed qed lemma binencode_decode: assumes "bit_symbols zs" and "even (length zs)" shows "binencode (bindecode zs) = zs" proof (rule nth_equalityI) let ?ys = "bindecode zs" show 1: "length (binencode ?ys) = length zs" using binencode_def bindecode_def assms(2) by fastforce show "binencode ?ys ! i = zs ! i" if "i < length (binencode ?ys)" for i proof - have ilen: "i < length zs" using 1 that by simp show ?thesis proof (cases "even i") case True let ?j = "i div 2" have jlen: "?j < length zs div 2" using ilen by (simp add: assms(2) less_mult_imp_div_less) then have ysj: "?ys ! ?j = decsym (zs ! (2 * ?j)) (zs ! (Suc (2 * ?j)))" using bindecode_def by simp have "binencode ?ys ! i = binencode ?ys ! (2 * ?j)" by (simp add: True) also have "... = 2 + (?ys ! ?j - 2) div 2" using binencode_at_even jlen by simp also have "... = zs ! (2 * ?j)" using ysj True assms(1) ilen by auto also have "... = zs ! i" using True by simp finally show "binencode ?ys ! i = zs ! i" . next case False let ?j = "i div 2" have jlen: "?j < length zs div 2" using ilen by (simp add: assms(2) less_mult_imp_div_less) then have ysj: "?ys ! ?j = decsym (zs ! (2 * ?j)) (zs ! (Suc (2 * ?j)))" using bindecode_def by simp have "binencode ?ys ! i = binencode ?ys ! (2 * ?j + 1)" by (simp add: False) also have "... = 2 + (?ys ! ?j - 2) mod 2" using binencode_at_odd jlen by simp also have "... = zs ! (2 * ?j + 1)" using ysj False assms(1) ilen by auto also have "... = zs ! i" using False by simp finally show "binencode ?ys ! i = zs ! i" . qed qed qed lemma binencode_inj: assumes "binencodable xs" and "binencodable ys" and "binencode xs = binencode ys" shows "xs = ys" using assms bindecode_encode by metis subsection \Turing machines for encoding and decoding\ text \ The next Turing machine implements @{const binencode} for @{const binencodable} symbol sequences. It expects a symbol sequence @{term zs} on tape $j_1$ and writes @{term "binencode zs"} to tape $j_2$. \ definition tm_binencode :: "tapeidx \ tapeidx \ machine" where "tm_binencode j1 j2 \ WHILE [] ; \rs. rs ! j1 \ \ DO IF \rs. rs ! j1 = \ THEN tm_print j2 [\, \] ELSE IF \rs. rs ! j1 = \ THEN tm_print j2 [\, \] ELSE IF \rs. rs ! j1 = \ THEN tm_print j2 [\, \] ELSE tm_print j2 [\, \] ENDIF ENDIF ENDIF ;; tm_right j1 DONE" lemma tm_binencode_tm: assumes "k \ 2" and "G \ 4" and "j1 < k" and "j2 < k" and "0 < j2" shows "turing_machine k G (tm_binencode j1 j2)" proof - have *: "symbols_lt G [\, \]" "symbols_lt G [\, \]" "symbols_lt G [\, \]" "symbols_lt G [\, \]" using assms(2) by (simp_all add: nth_Cons') then have "turing_machine k G (tm_print j2 [\, \])" using tm_print_tm[OF assms(5,4,2)] assms by blast moreover have "turing_machine k G (tm_print j2 [\, \])" using * tm_print_tm[OF assms(5,4,2)] assms by blast moreover have "turing_machine k G (tm_print j2 [\, \])" using * tm_print_tm[OF assms(5,4,2)] assms by blast moreover have "turing_machine k G (tm_print j2 [\, \])" using * tm_print_tm[OF assms(5,4,2)] assms by blast ultimately show ?thesis unfolding tm_binencode_def using turing_machine_loop_turing_machine Nil_tm turing_machine_branch_turing_machine tm_right_tm assms by simp qed locale turing_machine_binencode = fixes j1 j2 :: tapeidx begin definition "tm1 \ IF \rs. rs ! j1 = \ THEN tm_print j2 [\, \] ELSE tm_print j2 [\, \] ENDIF" definition "tm2 \ IF \rs. rs ! j1 = \ THEN tm_print j2 [\, \] ELSE tm1 ENDIF" definition "tm3 \ IF \rs. rs ! j1 = \ THEN tm_print j2 [\, \] ELSE tm2 ENDIF" definition "tm4 \ tm3 ;; tm_right j1" definition "tm5 \ WHILE [] ; \rs. rs ! j1 \ \ DO tm4 DONE" lemma tm5_eq_tm_binencode: "tm5 = tm_binencode j1 j2" using tm5_def tm4_def tm3_def tm2_def tm1_def tm_binencode_def by simp context fixes k :: nat and tps0 :: "tape list" and zs :: "symbol list" assumes jk: "k = length tps0" "j1 \ j2" "0 < j2" "j1 < k" "j2 < k" assumes zs: "binencodable zs" assumes tps0: "tps0 ! j1 = (\zs\, 1)" "tps0 ! j2 = (\[]\, 1)" begin definition tpsL :: "nat \ tape list" where "tpsL t \ tps0 [j1 := (\zs\, Suc t), j2 := (\binencode (take t zs)\, Suc (2 * t))]" lemma tpsL_0: "tpsL 0 = tps0" unfolding tpsL_def using tps0 jk by (metis One_nat_def length_0_conv length_binencode list_update_id mult_0_right take0) definition tpsL1 :: "nat \ tape list" where "tpsL1 t \ tps0 [j1 := (\zs\, Suc t), j2 := (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)]" lemma read_tpsL: assumes "t < length zs" shows "read (tpsL t) ! j1 = zs ! t" proof - have "tpsL t ! j1 = (\zs\, Suc t)" using tpsL_def jk by simp moreover have "j1 < length (tpsL t)" using tpsL_def jk by simp ultimately show "read (tpsL t) ! j1 = zs ! t" using assms tapes_at_read' by (metis Suc_leI contents_inbounds diff_Suc_1 fst_conv snd_conv zero_less_Suc) qed lemma tm1 [transforms_intros]: assumes "t < length zs" and "zs ! t = \ \ zs ! t = \" shows "transforms tm1 (tpsL t) 4 (tpsL1 t)" unfolding tm1_def proof (tform tps: jk tpsL_def) have 1: "tpsL t ! j2 = (\binencode (take t zs)\, Suc (2 * t))" using tpsL_def jk by simp have 2: "length (binencode (take t zs)) = 2 * t" using assms(1) by simp have "inscribe (tpsL t ! j2) [\, \] = (\binencode (take t zs) @ [\, \]\, Suc (2 * t) + 2)" using inscribe_contents 1 2 by (metis (no_types, lifting) One_nat_def Suc_1 Suc_eq_plus1 add_Suc_shift list.size(3) list.size(4)) moreover have "binencode (take t zs) @ [\, \] = binencode (take (Suc t) zs)" if "zs ! t = \" using that assms(1) binencode_snoc by (metis take_Suc_conv_app_nth) ultimately have 5: "inscribe (tpsL t ! j2) [\, \] = (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)" if "zs ! t = \" using that by simp have "inscribe (tpsL t ! j2) [\, \] = (\binencode (take t zs) @ [\, \]\, Suc (2 * t) + 2)" using inscribe_contents 1 2 by (metis (no_types, lifting) One_nat_def Suc_1 Suc_eq_plus1 add_Suc_shift list.size(3) list.size(4)) moreover have "binencode (take t zs) @ [\, \] = binencode (take (Suc t) zs)" if "zs ! t = 5" using that assms(1) binencode_snoc by (metis take_Suc_conv_app_nth) ultimately have 6: "inscribe (tpsL t ! j2) [\, \] = (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)" if "zs ! t = \" using that by simp have 7: "tpsL1 t = (tpsL t)[j2 := (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)]" unfolding tpsL1_def tpsL_def by (simp only: list_update_overwrite) then have 8: "tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" if "zs ! t = \" using that 5 by simp have 9: "tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" if "zs ! t = \" using that 6 7 by simp show "read (tpsL t) ! j1 = \ \ tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" using read_tpsL[OF assms(1)] 8 by simp show "read (tpsL t) ! j1 \ \ \ tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" using read_tpsL[OF assms(1)] 9 assms(2) by simp qed lemma tm2 [transforms_intros]: assumes "t < length zs" and "zs ! t = \ \ zs ! t = \ \ zs ! t = \" shows "transforms tm2 (tpsL t) 5 (tpsL1 t)" unfolding tm2_def proof (tform tps: tpsL_def jk assms(1)) have 1: "tpsL t ! j2 = (\binencode (take t zs)\, Suc (2 * t))" using tpsL_def jk by simp have 2: "length (binencode (take t zs)) = 2 * t" using assms(1) by simp show "read (tpsL t) ! j1 \ \ \ zs ! t = \ \ zs ! t = \" using read_tpsL[OF assms(1)] assms(2) by simp show "tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" if "read (tpsL t) ! j1 = \" proof - have *: "zs ! t = \" using read_tpsL[OF assms(1)] that by simp have "inscribe (tpsL t ! j2) [\, \] = (\binencode (take t zs) @ [2, \]\, Suc (2 * t) + 2)" using inscribe_contents 1 2 by (metis (no_types, lifting) One_nat_def Suc_1 Suc_eq_plus1 add_Suc_shift list.size(3) list.size(4)) moreover have "binencode (take t zs) @ [\, \] = binencode (take (Suc t) zs)" using * assms(1) binencode_snoc by (metis take_Suc_conv_app_nth) ultimately have "inscribe (tpsL t ! j2) [\, \] = (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)" by simp moreover have "tpsL1 t = (tpsL t)[j2 := (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)]" unfolding tpsL1_def tpsL_def by (simp only: list_update_overwrite) ultimately show "tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" using that by simp qed qed lemma tm3 [transforms_intros]: assumes "t < length zs" shows "transforms tm3 (tpsL t) 6 (tpsL1 t)" unfolding tm3_def proof (tform tps: jk assms tpsL_def) have 1: "tpsL t ! j2 = (\binencode (take t zs)\, Suc (2 * t))" using tpsL_def jk by simp have 2: "length (binencode (take t zs)) = 2 * t" using assms by simp show "read (tpsL t) ! j1 \ \ \ zs ! t = \ \ zs ! t = \ \ zs ! t = \" using assms zs read_tpsL by auto show "tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" if "read (tpsL t) ! j1 = \" proof - have *: "zs ! t = \" using read_tpsL[OF assms] that by simp have "inscribe (tpsL t ! j2) [\, \] = (\binencode (take t zs) @ [\, \]\, Suc (2 * t) + 2)" using inscribe_contents 1 2 by (metis (no_types, lifting) One_nat_def Suc_1 Suc_eq_plus1 add_Suc_shift list.size(3) list.size(4)) moreover have "binencode (take t zs) @ [\, \] = binencode (take (Suc t) zs)" using * assms binencode_snoc by (metis take_Suc_conv_app_nth) ultimately have "inscribe (tpsL t ! j2) [\, \] = (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)" by simp moreover have "tpsL1 t = (tpsL t)[j2 := (\binencode (take (Suc t) zs)\, Suc (2 * t) + 2)]" unfolding tpsL1_def tpsL_def by (simp only: list_update_overwrite) ultimately show "tpsL1 t = (tpsL t)[j2 := inscribe (tpsL t ! j2) [\, \]]" using that by simp qed qed lemma tm4 [transforms_intros]: assumes "t < length zs" shows "transforms tm4 (tpsL t) 7 (tpsL (Suc t))" unfolding tm4_def proof (tform tps: assms tpsL1_def jk) have *: "tpsL1 t ! j1 = (\zs\, Suc t)" using tpsL1_def jk by simp show "tpsL (Suc t) = (tpsL1 t)[j1 := tpsL1 t ! j1 |+| 1]" using tpsL_def tpsL1_def using jk * by (auto simp add: list_update_swap[of _ j1]) qed lemma tm5: assumes "ttt = 9 * length zs + 1" shows "transforms tm5 (tpsL 0) ttt (tpsL (length zs))" unfolding tm5_def proof (tform) show "\t. t < length zs \ read (tpsL t) ! j1 \ \" using read_tpsL zs by auto show "\ read (tpsL (length zs)) ! j1 \ \" proof - have "tpsL (length zs) ! j1 = (\zs\, Suc (length zs))" using tpsL_def jk by simp moreover have "j1 < length (tpsL (length zs))" using tpsL_def jk by simp ultimately have "read (tpsL (length zs)) ! j1 = tape_read (\zs\, Suc (length zs))" using tapes_at_read' by fastforce also have "... = \" using contents_outofbounds by simp finally show ?thesis by simp qed show "length zs * (7 + 2) + 1 \ ttt" using assms by simp qed lemma tpsL: "tpsL (length zs) = tps0 [j1 := (\zs\, Suc (length zs)), j2 := (\binencode zs\, Suc (2 * (length zs)))]" unfolding tpsL_def using tps0 jk by simp lemma tm5': assumes "ttt = 9 * length zs + 1" shows "transforms tm5 tps0 ttt (tpsL (length zs))" using assms tm5 tpsL_0 by simp end end (* locale turing_machine_binencode *) lemma transforms_tm_binencodeI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps tps' :: "tape list" and ttt k :: nat and zs :: "symbol list" assumes "k = length tps" "j1 \ j2" "0 < j2" "j1 < k" "j2 < k" and "binencodable zs" assumes "tps ! j1 = (\zs\, 1)" "tps ! j2 = (\[]\, 1)" assumes "ttt = 9 * length zs + 1" assumes "tps' \ tps [j1 := (\zs\, Suc (length zs)), j2 := (\binencode zs\, Suc (2 * (length zs)))]" shows "transforms (tm_binencode j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_binencode j1 j2 . show ?thesis using assms loc.tm5' loc.tm5_eq_tm_binencode loc.tpsL by simp qed text \ The next command reads chunks of two symbols over @{text "\\"} from one tape and writes to another tape the corresponding symbol over @{text "\\\\"}. The first symbol of each chunk is memorized on the last tape. If the end of the input (that is, a blank symbol) is encountered, the command stops without writing another symbol. \ definition cmd_bindec :: "tapeidx \ tapeidx \ command" where "cmd_bindec j1 j2 rs \ if rs ! j1 = 0 then (1, map (\z. (z, Stay)) rs) else (0, map (\i. if last rs = \ then if i = j1 then (rs ! i, Right) else if i = j2 then (rs ! i, Stay) else if i = length rs - 1 then (tosym (todigit (rs ! j1)), Stay) else (rs ! i, Stay) else if i = j1 then (rs ! i, Right) else if i = j2 then (decsym (last rs) (rs ! j1), Right) else if i = length rs - 1 then (1, Stay) else (rs ! i, Stay)) [0.. The Turing machine performing the decoding: \ definition tm_bindec :: "tapeidx \ tapeidx \ machine" where "tm_bindec j1 j2 = [cmd_bindec j1 j2]" context fixes j1 j2 :: tapeidx and k :: nat assumes j_less: "j1 < k" "j2 < k" and j_gt: "0 < j2" begin lemma turing_command_bindec: assumes "G \ 6" shows "turing_command (Suc k) 1 G (cmd_bindec j1 j2)" proof show "\gs. length gs = Suc k \ length ([!!] cmd_bindec j1 j2 gs) = length gs" using cmd_bindec_def by simp show "cmd_bindec j1 j2 gs [.] j < G" if "length gs = Suc k" "\i. i < length gs \ gs ! i < G" "j < length gs" for gs j proof (cases "gs ! j1 = \") case True then show ?thesis using that cmd_bindec_def by simp next case else: False show ?thesis proof (cases "last gs = \") case True then have "snd (cmd_bindec j1 j2 gs) = map (\i. if i = j1 then (gs ! i, Right) else if i = j2 then (gs ! i, Stay) else if i = length gs - 1 then (todigit (gs ! j1) + 2, Stay) else (gs ! i, Stay)) [0..i. if i = j1 then (gs ! i, Right) else if i = j2 then (2 * todigit (last gs) + todigit (gs ! j1) + 2, Right) else if i = length gs - 1 then (1, Stay) else (gs ! i, Stay)) [0.. length gs - 1" using that j_less by simp_all ultimately show ?thesis using cmd_bindec_def by simp next case False moreover have "0 < length gs" "0 \ j2" "0 \ length gs - 1" using that j_gt j_less by simp_all ultimately show ?thesis using cmd_bindec_def by simp qed show "\gs. length gs = Suc k \ [*] (cmd_bindec j1 j2 gs) \ 1" using cmd_bindec_def by simp qed lemma tm_bindec_tm: "G \ 6 \ turing_machine (Suc k) G (tm_bindec j1 j2)" unfolding tm_bindec_def using j_less(2) j_gt turing_command_bindec cmd_bindec_def by auto context fixes tps :: "tape list" and zs :: "symbol list" assumes j1_neq: "j1 \ j2" and lentps: "Suc k = length tps" and bs: "bit_symbols zs" begin lemma sem_cmd_bindec_gt: assumes "tps ! j1 = (\zs\, i)" and "i > length zs" shows "sem (cmd_bindec j1 j2) (0, tps) = (1, tps)" proof (rule semI) show "proper_command (Suc k) (cmd_bindec j1 j2)" using cmd_bindec_def by simp show "length tps = Suc k" using lentps by simp show "length tps = Suc k" using lentps by simp have "read tps ! j1 = \" using assms by (metis contents_outofbounds fst_conv j_less(1) lentps less_Suc_eq snd_conv tapes_at_read') moreover from this show "fst (cmd_bindec j1 j2 (read tps)) = 1" by (simp add: cmd_bindec_def) ultimately show "\j. j < Suc k \ act (cmd_bindec j1 j2 (read tps) [!] j) (tps ! j) = tps ! j" using assms cmd_bindec_def act_Stay lentps read_length by simp qed lemma sem_cmd_bindec_1: assumes "tps ! k = \\\" and "tps ! j1 = (\zs\, i)" and "i > 0" and "i \ length zs" and "tps' = tps [j1 := tps ! j1 |+| 1, k := \todigit (tps :.: j1) + 2\]" shows "sem (cmd_bindec j1 j2) (0, tps) = (0, tps')" proof (rule semI) show "proper_command (Suc k) (cmd_bindec j1 j2)" using cmd_bindec_def by simp show "length tps = Suc k" using lentps by simp show "length tps' = Suc k" using lentps assms(5) by simp have read: "read tps ! j1 \ \" using assms(2,3,4) bs j_less(1) tapes_at_read'[of j1 tps] contents_inbounds[OF assms(3,4)] lentps proper_symbols_ne0[OF proper_bit_symbols[OF bs]] by (metis One_nat_def Suc_less_eq Suc_pred fst_conv le_imp_less_Suc less_SucI snd_eqD) then show "fst (cmd_bindec j1 j2 (read tps)) = 0" using cmd_bindec_def by simp show "act (cmd_bindec j1 j2 (read tps) [!] j) (tps ! j) = tps' ! j" if "j < Suc k" for j proof - let ?rs = "read tps" have "last ?rs = 1" using assms(1) lentps onesie_read read_length tapes_at_read' by (metis (mono_tags, lifting) last_length lessI) then have *: "snd (cmd_bindec j1 j2 ?rs) = map (\i. if i = j1 then (?rs ! i, Right) else if i = j2 then (?rs ! i, Stay) else if i = length ?rs - 1 then (if ?rs ! j1 = \ then \ else \, Stay) else (?rs ! i, Stay)) [0..length ?rs = Suc k\) consider "j = j1" | "j \ j1 \ j = j2" | "j \ j1 \ j \ j2 \ j = k" | "j \ j1 \ j \ j2 \ j \ k" by auto then show ?thesis proof (cases) case 1 then have "cmd_bindec j1 j2 ?rs [!] j = (?rs ! j1, Right)" using * len by simp then show ?thesis using act_Right 1 assms(5) j_less(1) lentps by simp next case 2 then have "cmd_bindec j1 j2 ?rs [!] j = (?rs ! j2, Stay)" using * len by simp then show ?thesis using 2 act_Stay assms(5) j_less(2) lentps by simp next case 3 then have "cmd_bindec j1 j2 ?rs [!] j = (todigit (?rs ! j1) + 2, Stay)" using k * len by simp then show ?thesis using 3 assms(1,5) act_onesie j_less(1) lentps tapes_at_read' by (metis length_list_update less_Suc_eq nth_list_update) next case 4 then have "cmd_bindec j1 j2 ?rs [!] j = (?rs ! j, Stay)" using k * len that by simp then show ?thesis using 4 act_Stay assms(5) lentps that by simp qed qed qed lemma sem_cmd_bindec_23: assumes "tps ! k = \s\" and "s = \ \ s = \" and "tps ! j1 = (\zs\, i)" and "i > 0" and "i \ length zs" and "tps' = tps [j1 := tps ! j1 |+| 1, j2 := tps ! j2 |:=| decsym s (tps :.: j1) |+| 1, k := \\\]" shows "sem (cmd_bindec j1 j2) (0, tps) = (0, tps')" proof (rule semI) show "proper_command (Suc k) (cmd_bindec j1 j2)" using cmd_bindec_def by simp show "length tps = Suc k" using lentps by simp show "length tps' = Suc k" using lentps assms(6) by simp have read: "read tps ! j1 \ \" using assms(3-5) bs tapes_at_read'[of j1 tps] contents_inbounds[OF assms(4,5)] lentps by (metis One_nat_def Suc_less_eq Suc_pred fst_conv j_less(1) le_imp_less_Suc less_imp_le_nat not_one_less_zero proper_bit_symbols snd_conv) show "fst (cmd_bindec j1 j2 (read tps)) = 0" using read cmd_bindec_def by simp show "act (cmd_bindec j1 j2 (read tps) [!] j) (tps ! j) = tps' ! j" if "j < Suc k" for j proof - let ?rs = "read tps" have "last ?rs \ 1" using assms(1,2) lentps onesie_read read_length tapes_at_read' by (metis Suc_1 diff_Suc_1 last_conv_nth lessI list.size(3) n_not_Suc_n numeral_One numeral_eq_iff old.nat.distinct(1) semiring_norm(86)) then have *: "snd (cmd_bindec j1 j2 ?rs) = map (\i. if i = j1 then (?rs ! i, Right) else if i = j2 then (2 * todigit (last ?rs) + todigit (?rs ! j1) + 2, Right) else if i = length ?rs - 1 then (1, Stay) else (?rs ! i, Stay)) [0.. j1 \ j = j2" | "j \ j1 \ j \ j2 \ j = k" | "j \ j1 \ j \ j2 \ j \ k" by auto then show ?thesis proof (cases) case 1 then show ?thesis using * len act_Right assms(6) j_less(1) j1_neq lentps by simp next case 2 then have "cmd_bindec j1 j2 ?rs [!] j = (2 * todigit (last ?rs) + todigit (?rs ! j1) + 2, Right)" using * len by simp moreover have "last ?rs = s" using assms(1,2) lenrs k onesie_read tapes_at_read' by (metis last_conv_nth length_0_conv lentps lessI old.nat.distinct(1)) moreover have "?rs ! j1 = tps :.: j1" using j_less(1) lentps tapes_at_read' by simp ultimately show ?thesis using 2 assms(6) act_Right' j_less lentps by simp next case 3 then show ?thesis using * len k act_onesie assms(1,6) lentps by simp next case 4 then have "cmd_bindec j1 j2 ?rs [!] j = (?rs ! j, Stay)" using k * len by simp then show ?thesis using 4 act_Stay assms(6) lentps that by simp qed qed qed end (* tps zs *) lemma transits_tm_bindec: fixes tps :: "tape list" and zs :: "symbol list" assumes j1_neq: "j1 \ j2" and j1j2: "0 < j2" "j1 < k" "j2 < k" and lentps: "Suc k = length tps" and bs: "bit_symbols zs" assumes "tps ! k = \\\" and "tps ! j1 = (\zs\, 2 * i + 1)" and "tps ! j2 = (\bindecode (take (2 * i) zs)\, Suc i)" and "i < length zs div 2" and "tps' = tps [j1 := (\zs\, 2 * (Suc i) + 1), j2 := (\bindecode (take (2 * Suc i) zs)\, Suc (Suc i))]" shows "transits (tm_bindec j1 j2) (0, tps) 2 (0, tps')" proof - define tps1 where "tps1 = tps [j1 := (\zs\, 2 * i + 2), k := \todigit (tps :.: j1) + 2\]" let ?i = "2 * i + 1" let ?M = "tm_bindec j1 j2" have ilen: "?i < length zs" using assms(10) by simp have "exe ?M (0, tps) = sem (cmd_bindec j1 j2) (0, tps)" using tm_bindec_def exe_lt_length by simp also have "... = (if ?i \ length zs then 0 else 1, tps[j1 := tps ! j1 |+| 1, k := \todigit (tps :.: j1) + 2\])" using ilen bs assms(7,8) sem_cmd_bindec_1 j1_neq lentps by simp also have "... = (0, tps[j1 := tps ! j1 |+| 1, k := \todigit (tps :.: j1) + 2\])" using ilen by simp also have "... = (0, tps1)" using tps1_def assms by simp finally have step1: "exe ?M (0, tps) = (0, tps1)" . let ?s = "tps1 :.: k" have "tps :.: j1 = zs ! (2 * i)" using assms(8) ilen by simp then have "?s = todigit (zs ! (2 * i)) + 2" using tps1_def lentps by simp then have "?s = zs ! (2 * i)" - using ilen bs by (smt (z3) One_nat_def Suc_1 add_2_eq_Suc' add_lessD1 numeral_3_eq_3) + using ilen bs by (smt (verit) One_nat_def Suc_1 add_2_eq_Suc' add_lessD1 numeral_3_eq_3) moreover have "tps1 :.: j1 = zs ! ?i" using tps1_def ilen lentps j1j2 by simp ultimately have *: "decsym ?s (tps1 :.: j1) = decsym (zs ! (2*i)) (zs ! (Suc (2*i)))" by simp have "exe ?M (0, tps1) = sem (cmd_bindec j1 j2) (0, tps1)" using tm_bindec_def exe_lt_length by simp also have "... = (if Suc ?i \ length zs then 0 else 1, tps1[j1 := tps1 ! j1 |+| 1, j2 := tps1 ! j2 |:=| 2 * todigit ?s + todigit (tps1 :.: j1) + 2 |+| 1, k := \\\])" proof - have 1: "tps1 ! k = \?s\" using tps1_def lentps by simp have 2: "?s = 2 \ ?s = 3" using tps1_def lentps by simp have 3: "tps1 ! j1 = (\zs\, Suc ?i)" using tps1_def lentps Suc_1 add_Suc_right j_less(1) less_Suc_eq nat_neq_iff nth_list_update_eq nth_list_update_neq by simp have 4: "Suc ?i > 0" by simp have 5: "Suc k = length tps1" by (simp add: lentps tps1_def) show ?thesis using ilen sem_cmd_bindec_23[of tps1 zs ?s "Suc ?i", OF j1_neq 5 bs 1 2 3 4] by simp qed also have "... = (0, tps1[j1 := tps1 ! j1 |+| 1, j2 := tps1 ! j2 |:=| decsym ?s (tps1 :.: j1) |+| 1, k := \\\])" using assms(10) length_binencode by simp also have "... = (0, tps1[j1 := tps1 ! j1 |+| 1, j2 := tps1 ! j2 |:=| decsym (zs ! (2*i)) (zs ! (Suc (2*i))) |+| 1, k := \\\])" (is "_ = (0, ?tps)") using * by simp also have "... = (0, tps')" proof - have len': "length tps' = Suc k" using assms lentps by simp have len1: "length tps1 = Suc k" using assms lentps tps1_def by simp have 1: "?tps ! k = tps' ! k" using assms(7,11) by (simp add: j_less(1) j_less(2) len1 nat_neq_iff) have 2: "?tps ! j1 = tps' ! j1" using assms(11) j1_neq j_less(1) lentps tps1_def by simp have "?tps ! j2 = tps1 ! j2 |:=| decsym (zs ! (2*i)) (zs ! (Suc (2*i))) |+| 1" by (simp add: j_less(2) len1 less_Suc_eq nat_neq_iff) also have "... = tps ! j2 |:=| decsym (zs ! (2*i)) (zs ! (Suc (2*i))) |+| 1" using tps1_def j1_neq j_less(2) len1 by force also have "... = (\bindecode (take (2 * i) zs)\, Suc i) |:=| decsym (zs ! (2*i)) (zs ! (Suc (2*i))) |+| 1" using assms(9) j1_neq j_less(2) len1 tps1_def by simp also have "... = (\bindecode (take (2 * i) zs)\(Suc i := decsym (zs ! (2*i)) (zs ! (Suc (2*i)))), Suc (Suc i))" by simp also have "... = (\bindecode (take (2 * i) zs) @ [decsym (zs ! (2*i)) (zs ! (Suc (2*i)))]\, Suc (Suc i))" using contents_snoc[of "bindecode (take (2 * i) zs)"] ilen length_bindecode proof - have "length (bindecode (take (2 * i) zs)) = i" using ilen length_bindecode by simp then show ?thesis using contents_snoc[of "bindecode (take (2 * i) zs)"] by simp qed also have "... = (\bindecode (take (2 * Suc i) zs)\, Suc (Suc i))" using bindecode_take_snoc ilen by simp - also have "... = tps' ! j2" + also have *: "... = tps' ! j2" by (metis assms(11) j_less(2) length_list_update lentps less_Suc_eq nth_list_update_eq) finally have "?tps ! j2 = tps' ! j2" . - then show ?thesis - using 1 2 assms(11) tps1_def by (smt (z3) list_update_id list_update_overwrite list_update_swap) + with 1 2 assms(11) * show ?thesis + unfolding tps1_def + by (smt (verit) j_less(1) j_less(2) lentps less_Suc_eq list_update_id list_update_overwrite list_update_swap nat_neq_iff nth_list_update_eq nth_list_update_neq) qed finally have "exe ?M (0, tps1) = (0, tps')" . then have "execute ?M (0, tps) 2 = (0, tps')" using step1 by (simp add: numeral_2_eq_2) then show "transits (tm_bindec j1 j2) (0, tps) 2 (0, tps')" using execute_imp_transits by blast qed context fixes tps :: "tape list" and zs :: "symbol list" assumes j1_neq: "j1 \ j2" and j1j2: "0 < j2" "j1 < k" "j2 < k" and lentps: "Suc k = length tps" and bs: "bit_symbols zs" begin lemma transits_tm_bindec': assumes "tps ! k = \\\" and "tps ! j1 = (\zs\, 1)" and "tps ! j2 = (\[]\, 1)" and "i \ length zs div 2" and "tps' = tps [j1 := (\zs\, 2 * i + 1), j2 := (\bindecode (take (2 * i) zs)\, Suc i)]" shows "transits (tm_bindec j1 j2) (0, tps) (2 * i) (0, tps')" using assms(4,5) proof (induction i arbitrary: tps') case 0 then show ?case using assms(2,3) by (metis One_nat_def add.commute div_mult_self1_is_m execute.simps(1) le_numeral_extra(3) length_bindecode length_greater_0_conv list.size(3) list_update_id mult_0_right plus_1_eq_Suc take0 transits_def zero_less_numeral) next case (Suc i) define tpsi where "tpsi = tps [j1 := (\zs\, 2 * i + 1), j2 := (\bindecode (take (2*i) zs)\, Suc i)]" then have "transits (tm_bindec j1 j2) (0, tps) (2 * i) (0, tpsi)" using Suc by simp moreover have "transits (tm_bindec j1 j2) (0, tpsi) 2 (0, tps')" proof - have 1: "tpsi ! k = \\\" using tpsi_def by (simp add: assms(1) j_less(1) j_less(2) less_not_refl3) have 2: "tpsi ! j1 = (\zs\, 2 * i + 1)" using tpsi_def by (metis j1_neq j_less(1) lentps less_Suc_eq nth_list_update_eq nth_list_update_neq) have 3: "tpsi ! j2 = (\bindecode (take (2 * i) zs)\, Suc i)" using tpsi_def by (metis j_less(2) length_list_update lentps less_Suc_eq nth_list_update_eq) have 4: "i < length zs div 2" using Suc by simp have 5: "tps' = tpsi [j1 := (\zs\, 2 * (Suc i) + 1), j2 := (\bindecode (take (2 * Suc i) zs)\, Suc (Suc i))]" using Suc tpsi_def by (metis (no_types, opaque_lifting) list_update_overwrite list_update_swap) have 6: "Suc k = length tpsi" using tpsi_def lentps by simp show ?thesis using transits_tm_bindec[OF j1_neq j1j2 6 bs 1 2 3 4 5] . qed ultimately show "transits (tm_bindec j1 j2) (0, tps) (2 * (Suc i)) (0, tps')" using transits_additive by fastforce qed corollary transits_tm_bindec'': assumes "tps ! k = \\\" and "tps ! j1 = (\zs\, 1)" and "tps ! j2 = (\[]\, 1)" and "l = length zs div 2" and "tps' = tps [j1 := (\zs\, 2 * l + 1), j2 := (\bindecode (take (2 * l) zs)\, Suc l)]" shows "transits (tm_bindec j1 j2) (0, tps) (2 * l) (0, tps')" using assms transits_tm_bindec' by simp text \In case the input is of odd length, that is, malformed:\ lemma transforms_tm_bindec_odd: assumes "tps ! k = \\\" and "tps ! j1 = (\zs\, 1)" and "tps ! j2 = (\[]\, 1)" and "tps' = tps [j1 := (\zs\, 2 * l + 2), j2 := (\bindecode zs\, Suc l), k := \todigit (last zs) + 2\]" and "l = length zs div 2" and "Suc (2 * l) = length zs" shows "transforms (tm_bindec j1 j2) tps (2 * l + 2) tps'" proof - let ?ys = "bindecode (take (2 * l) zs)" let ?i = "2 * l + 1" let ?M = "tm_bindec j1 j2" have ys: "?ys = bindecode zs" using bindecode_odd assms(6) by (metis Suc_eq_plus1) have "zs \ []" using assms(6) by auto define tps1 where "tps1 = tps [j1 := (\zs\, 2 * l + 1), j2 := (\?ys\, Suc l)]" define tps2 where "tps2 = tps [j1 := (\zs\, 2 * l + 2), j2 := (\bindecode zs\, Suc l), k := \todigit (tps1 :.: j1) + 2\]" have "transits ?M (0, tps) (2 * l) (0, tps1)" using tps1_def assms transits_tm_bindec'' by simp moreover have "execute ?M (0, tps1) 1 = (0, tps2)" proof - have "execute ?M (0, tps1) 1 = exe ?M (0, tps1)" by simp also have "... = sem (cmd_bindec j1 j2) (0, tps1)" using exe_lt_length tm_bindec_def by simp also have "... = (0, tps1[j1 := tps1 ! j1 |+| 1, k := \todigit (tps1 :.: j1) + 2\])" (is "_ = (0, ?tps)") proof - have "tps1 ! j1 = (\zs\, ?i)" using lentps tps1_def j1_neq j_less by simp moreover have "?i > 0" by simp moreover have "tps1 ! k = tps ! k" using tps1_def by (simp add: j_less(1) j_less(2) nat_neq_iff) moreover have "?i \ length zs" by (simp add: assms(6)) ultimately have "sem (cmd_bindec j1 j2) (0, tps1) = (0, ?tps)" using sem_cmd_bindec_1 assms(1,4) bit_symbols_binencode bs j1_neq lentps tps1_def by (metis length_list_update) then show ?thesis by simp qed also have "... = (0, tps2)" proof - have "tps2 ! j1 = ?tps ! j1" using tps1_def tps2_def j1_neq j_less(1) lentps by simp moreover have "tps2 ! j2 = ?tps ! j2" using tps1_def tps2_def j1_neq j_less(2) lentps ys by simp ultimately have "tps2 = ?tps" using tps2_def tps1_def j_less(1) lentps - by (smt (z3) list_update_id list_update_overwrite list_update_swap) + by (smt (verit) list_update_id list_update_overwrite list_update_swap) then show ?thesis by simp qed finally show ?thesis . qed ultimately have "transits ?M (0, tps) (2 * l + 1) (0, tps2)" using execute_imp_transits transits_additive by blast moreover have "execute ?M (0, tps2) 1 = (1, tps')" proof - have "execute ?M (0, tps2) 1 = exe ?M (0, tps2)" by simp also have "... = sem (cmd_bindec j1 j2) (0, tps2)" using exe_lt_length tm_bindec_def by simp also have "... = (1, tps2)" proof - have "2 * l + 2 > length zs" using assms(5,6) by simp moreover have "tps2 ! j1 = (\zs\, 2 * l + 2)" using tps2_def j1_neq j_less(1) lentps by simp ultimately show ?thesis using sem_cmd_bindec_gt[of tps2 zs "2 * l + 2"] by (metis bs j1_neq length_list_update lentps tps2_def) qed moreover have "tps2 = tps'" proof - have "tps1 :.: j1 = last zs" using tps1_def assms \zs \ []\ contents_inbounds by (metis Suc_leI add.commute fst_conv j1_neq j_less(1) last_conv_nth lentps less_Suc_eq nth_list_update_eq nth_list_update_neq plus_1_eq_Suc snd_conv zero_less_Suc) then show ?thesis using tps2_def assms(4) by simp qed ultimately show ?thesis by simp qed ultimately have "transits ?M (0, tps) (2 * l + 2) (1, tps')" - using execute_imp_transits transits_additive by (smt (z3) ab_semigroup_add_class.add_ac(1) nat_1_add_1) + using execute_imp_transits transits_additive by (smt (verit) ab_semigroup_add_class.add_ac(1) nat_1_add_1) then show "transforms (tm_bindec j1 j2) tps (2 * l + 2) tps'" using transforms_def tm_bindec_def by simp qed text \ In case the input is of even length, that is, properly encoded: \ lemma transforms_tm_bindec_even: assumes "tps ! k = \\\" and "tps ! j1 = (\zs\, 1)" and "tps ! j2 = (\[]\, 1)" and "tps' = tps [j1 := (\zs\, 2 * l + 1), j2 := (\bindecode zs\, Suc l)]" and "l = length zs div 2" and "2 * l = length zs" shows "transforms (tm_bindec j1 j2) tps (2 * l + 1) tps'" proof - let ?ys = "bindecode (take (2 * l) zs)" let ?i = "2 * l + 1" let ?M = "tm_bindec j1 j2" have ys: "?ys = bindecode zs" using assms(6) by simp have "transits ?M (0, tps) (2 * l) (0, tps')" using assms transits_tm_bindec'' by simp moreover have "execute ?M (0, tps') 1 = (1, tps')" proof - have "execute ?M (0, tps') 1 = exe ?M (0, tps')" using assms(4) by simp also have "... = sem (cmd_bindec j1 j2) (0, tps')" using exe_lt_length tm_bindec_def by simp also have "... = (1, tps')" proof - have "tps' ! j1 = (\zs\, ?i)" using lentps assms(4) j1_neq j_less by simp moreover have "?i > length zs" using assms(6) by simp moreover have "tps' ! k = tps ! k" using assms(4) by (simp add: j_less(1) j_less(2) nat_neq_iff) ultimately have "sem (cmd_bindec j1 j2) (0, tps') = (1, tps')" using sem_cmd_bindec_gt assms(1,4) bit_symbols_binencode bs j1_neq lentps assms(4) by simp then show ?thesis by simp qed finally show ?thesis . qed ultimately have "transits ?M (0, tps) (2 * l + 1) (1, tps')" using execute_imp_transits transits_additive by blast then show "transforms (tm_bindec j1 j2) tps (2 * l + 1) tps'" using tm_bindec_def transforms_def by simp qed lemma transforms_tm_bindec: assumes "tps ! k = \\\" and "tps ! j1 = (\zs\, 1)" and "tps ! j2 = (\[]\, 1)" and "tps' = tps [j1 := (\zs\, Suc (length zs)), j2 := (\bindecode zs\, Suc (length zs div 2)), k := \if even (length zs) then 1 else (todigit (last zs) + 2)\]" shows "transforms (tm_bindec j1 j2) tps (Suc (length zs)) tps'" proof (cases "even (length zs)") case True then show ?thesis using transforms_tm_bindec_even[OF assms(1-3)] assms(1,4) j_less(1) j_less(2) - by (smt (z3) Suc_eq_plus1 dvd_mult_div_cancel list_update_id list_update_swap nat_neq_iff) + by (smt (verit) Suc_eq_plus1 dvd_mult_div_cancel list_update_id list_update_swap nat_neq_iff) next case False then show ?thesis using assms(4) transforms_tm_bindec_odd[OF assms(1-3)] by simp qed end (* context tps zs *) end (* context j1 j2 k *) text \ Next we eliminate the memorization tape from @{const tm_bindec}. \ lemma transforms_cartesian_bindec: assumes "G \ (6 :: nat)" assumes "j1 \ j2" and j1j2: "0 < j2" "j1 < k" "j2 < k" and "k = length tps" and "bit_symbols zs" assumes "tps ! j1 = (\zs\, 1)" and "tps ! j2 = (\[]\, 1)" assumes "t = Suc (length zs)" and "tps' = tps [j1 := (\zs\, Suc (length zs)), j2 := (\bindecode zs\, Suc (length zs div 2))]" shows "transforms (cartesian (tm_bindec j1 j2) 4) tps t tps'" proof (rule cartesian_transforms_onesie) show "turing_machine (Suc k) G (tm_bindec j1 j2)" using tm_bindec_tm assms(1) j1j2 by simp show "immobile (tm_bindec j1 j2) k (Suc k)" proof (standard+) fix q :: nat and rs :: "symbol list" assume "q < length (tm_bindec j1 j2)" "length rs = Suc k" then have *: "tm_bindec j1 j2 ! q = cmd_bindec j1 j2" using tm_bindec_def by simp moreover have "cmd_bindec j1 j2 rs [~] k = Stay" using cmd_bindec_def `length rs = Suc k` j1j2 by (smt (verit, best) add_diff_inverse_nat diff_zero length_upt lessI less_nat_zero_code nat_neq_iff nth_map nth_upt prod.sel(2)) ultimately show "(tm_bindec j1 j2 ! q) rs [~] k = Stay" using * by simp qed show "2 \ k" using j1j2 by linarith show "(1::nat) < 4" by simp show "length tps = k" using assms(3,6) by simp show "bounded_write (tm_bindec j1 j2) k 4" proof - { fix q :: nat and rs :: "symbol list" assume q: "q < length (tm_bindec j1 j2)" and rs: "length rs = Suc k" then have "tm_bindec j1 j2 ! q = cmd_bindec j1 j2" using tm_bindec_def by simp have "cmd_bindec j1 j2 rs [.] (length rs - 1) < 4 \ fst (cmd_bindec j1 j2 rs) = 1" proof (cases "rs ! j1 = 0") case True then show ?thesis using cmd_bindec_def by simp next case else: False show ?thesis proof (cases "last rs = 1") case True then have "snd (cmd_bindec j1 j2 rs) = map (\i. if i = j1 then (rs ! i, Right) else if i = j2 then (rs ! i, Stay) else if i = length rs - 1 then (todigit (rs ! j1) + 2, Stay) else (rs ! i, Stay)) [0..i. if i = j1 then (rs ! i, Right) else if i = j2 then (2 * todigit (last rs) + todigit (rs ! j1) + 2, Right) else if i = length rs - 1 then (1, Stay) else (rs ! i, Stay)) [0..\\]) t (tps' @ [\?c\])" (is "transforms _ ?tps t ?tps'") proof - have "?tps ! k = \\\" by (simp add: assms(6)) moreover have "?tps ! j1 = (\zs\, 1)" by (metis assms(6) assms(8) j1j2(2) nth_append) moreover have "?tps ! j2 = (\[]\, 1)" by (metis assms(6) assms(9) j1j2(3) nth_append) moreover have "?tps' = ?tps [j1 := (\zs\, Suc (length zs)), j2 := (\bindecode zs\, Suc (length zs div 2)), k := \?c\]" by (metis (no_types, lifting) assms(6,11) j1j2(2,3) length_list_update list_update_append1 list_update_length) ultimately show ?thesis using transforms_tm_bindec[of j1 "k" j2 ?tps zs ?tps'] assms by simp qed qed text \ The next Turing machine decodes a bit symbol sequence given on tape $j_1$ into a quaternary symbol sequence output to tape $j_2$. It executes the previous TM followed by carriage returns on the tapes $j_1$ and $j_2$. \ definition tm_bindecode :: "tapeidx \ tapeidx \ machine" where "tm_bindecode j1 j2 \ cartesian (tm_bindec j1 j2) 4 ;; tm_cr j1 ;; tm_cr j2" lemma tm_bindecode_tm: fixes j1 j2 :: tapeidx and G k :: nat assumes "G \ 6" and "j1 < k" and "j2 < k" and "0 < j2" and "j1 \ j2" shows "turing_machine k G (tm_bindecode j1 j2)" using assms tm_bindec_tm tm_bindecode_def cartesian_tm tm_cr_tm by simp locale turing_machine_bindecode = fixes j1 j2 :: tapeidx begin definition "tm1 \ cartesian (tm_bindec j1 j2) 4" definition "tm2 \ tm1 ;; tm_cr j1" definition "tm3 \ tm2 ;; tm_cr j2" lemma tm3_eq_tm_bindecode: "tm3 = tm_bindecode j1 j2" using tm1_def tm2_def tm3_def tm_bindecode_def by simp context fixes tps0 :: "tape list" and zs :: "symbol list" and k :: nat assumes jk: "j1 < k" "j2 < k" "0 < j2" "j1 \ j2" "k = length tps0" assumes zs: "bit_symbols zs" assumes tps0: "tps0 ! j1 = (\zs\, 1)" "tps0 ! j2 = (\[]\, 1)" begin definition "tps1 \ tps0 [j1 := (\zs\, Suc (length zs)), j2 := (\bindecode zs\, Suc (length zs div 2))]" lemma tm1 [transforms_intros]: assumes "t = Suc (length zs)" shows "transforms tm1 tps0 t tps1" unfolding tm1_def using transforms_cartesian_bindec assms jk tps0 zs tps1_def by blast definition "tps2 \ tps0 [j2 := (\bindecode zs\, Suc (length zs div 2))]" lemma tm2 [transforms_intros]: assumes "t = 2 * length zs + 4" shows "transforms tm2 tps0 t tps2" unfolding tm2_def proof (tform tps: assms) show "j1 < length tps1" using jk tps1_def by simp show "clean_tape (tps1 ! j1)" using jk zs clean_contents_proper tps1_def by fastforce show "tps2 = tps1[j1 := tps1 ! j1 |#=| 1]" using tps0 jk tps2_def tps1_def by (metis (no_types, lifting) fst_conv list_update_id list_update_overwrite list_update_swap nth_list_update_eq nth_list_update_neq) show "t = Suc (length zs) + (tps1 :#: j1 + 2)" using assms(1) jk tps1_def by simp qed definition "tps3 \ tps0 [j2 := (\bindecode zs\, 1)]" lemma tm3: assumes "t = 2 * length zs + 7 + length zs div 2" shows "transforms tm3 tps0 t tps3" unfolding tm3_def proof (tform tps: jk tps2_def tps3_def assms) show "clean_tape (tps2 ! j2)" using jk bindecode_at tps2_def by simp qed lemma tm3': assumes "t = 7 + 3 * length zs" shows "transforms tm3 tps0 t tps3" proof - have "7 + 3 * length zs \ 2 * length zs + 7 + length zs div 2" by simp then show ?thesis using transforms_monotone tm3 assms tps3_def by simp qed end (* context tps zs *) end (* locale turing_machine_bindecode *) lemma transforms_tm_bindecodeI [transforms_intros]: fixes j1 j2 :: tapeidx fixes tps :: "tape list" and zs :: "symbol list" and k ttt :: nat assumes "j1 < k" and "j2 < k" and "0 < j2" and "j1 \ j2" and "k = length tps" and "bit_symbols zs" assumes "tps ! j1 = (\zs\, 1)" "tps ! j2 = (\[]\, 1)" assumes "ttt = 7 + 3 * length zs" assumes "tps' = tps [j2 := (\bindecode zs\, 1)]" shows "transforms (tm_bindecode j1 j2) tps ttt tps'" proof - interpret loc: turing_machine_bindecode j1 j2 . show ?thesis using loc.tm3_eq_tm_bindecode loc.tm3' loc.tps3_def assms by simp qed end